diff --git a/src/Init/Data/Option/Basic.lean b/src/Init/Data/Option/Basic.lean index 8fa868d61200..c36b79fb9639 100644 --- a/src/Init/Data/Option/Basic.lean +++ b/src/Init/Data/Option/Basic.lean @@ -292,6 +292,12 @@ instance (r : α → β → Prop) [s : DecidableRel r] : DecidableRel (Option.lt | some _, none => isFalse not_false | none, none => isFalse not_false +instance (r : α → β → Prop) [s : DecidableRel r] : DecidableRel (Option.le r) + | none, some _ => isTrue trivial + | some x, some y => s x y + | some _, none => isFalse not_false + | none, none => isTrue trivial + namespace SomeLtNone /-- diff --git a/src/Lean/Data.lean b/src/Lean/Data.lean index 543f9e699c20..467f0f29ae5e 100644 --- a/src/Lean/Data.lean +++ b/src/Lean/Data.lean @@ -30,3 +30,4 @@ public import Lean.Data.NameTrie public import Lean.Data.RBTree public import Lean.Data.RBMap public import Lean.Data.RArray +public import Lean.Data.Fmt diff --git a/src/Lean/Data/Fmt.lean b/src/Lean/Data/Fmt.lean new file mode 100644 index 000000000000..8d206fb4ac39 --- /dev/null +++ b/src/Lean/Data/Fmt.lean @@ -0,0 +1,13 @@ +/- +Copyright (c) 2025 Lean FRO. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. + +Author: Marc Huisinga +-/ +module + +prelude +public import Lean.Data.Fmt.Basic +public import Lean.Data.Fmt.Formatter +public import Lean.Data.Fmt.LawfulCost +public import Lean.Data.Fmt.Json diff --git a/src/Lean/Data/Fmt/Basic.lean b/src/Lean/Data/Fmt/Basic.lean new file mode 100644 index 000000000000..3dc01eeba6b6 --- /dev/null +++ b/src/Lean/Data/Fmt/Basic.lean @@ -0,0 +1,664 @@ +/- +Copyright (c) 2025 Lean FRO. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. + +Author: Marc Huisinga +-/ +module + +prelude +public import Init.Data.Hashable +import Init.Data.Array + +/-! +Document language of the `Fmt` formatter. + +This file implements the document language of 'A Pretty Expressive Printer' [1] by +Sorawee Porncharoenwase, Justin Pombrio and Emina Torlak. +This implementation is based on the Racket implementation of pretty-expressive [2]. + +[1] https://arxiv.org/pdf/2310.01530 +[2] https://docs.racket-lang.org/pretty-expressive/ +-/ + +public section + +namespace Lean.Fmt + +/-- +Document height at which the formatting document is memoized when formatting. +Time-complexity-wise, this is sound because the formatting document is a binary tree, and so its +height bounds the amount of nodes in the tree. +-/ +def memoHeightLimit : Nat := 6 + +/-- +Computes the memoization height of a node in the formatting document with a child memoization height +of `memoHeight`. +-/ +def nextMemoHeight (memoHeight : Nat) : Nat := + if memoHeight = 0 then + memoHeightLimit + else + memoHeight - 1 + +/-- +Bitmap that tracks whether there is a `Doc.full` node on the same line *before* the current document +that is being resolved by the formatter or on the same line *after* the current document. + +In the formatter, we case split on the fullness state in several places and then prune subtrees +of the search when we notice that they are inconsistent with the actual document currently being +resolved. +-/ +@[expose] +def FullnessState := UInt8 + deriving BEq, Hashable + +@[inline] +def FullnessState.mk (isFullBefore : Bool) (isFullAfter : Bool) : FullnessState := + (isFullBefore.toUInt8 <<< 1) ||| isFullAfter.toUInt8 + +@[inline] +def FullnessState.isFullBefore (s : FullnessState) : Bool := + let s : UInt8 := s + (s &&& 0b10) == 1 + +@[inline] +def FullnessState.isFullAfter (s : FullnessState) : Bool := + let s : UInt8 := s + (s &&& 0b1) == 1 + +@[inline] +def FullnessState.setFullBefore (s : FullnessState) (isFullBefore : Bool) : FullnessState := + let s : UInt8 := s + (s &&& (0b11111101 : UInt8)) ||| (isFullBefore.toUInt8 <<< 1) + +@[inline] +def FullnessState.setFullAfter (s : FullnessState) (isFullAfter : Bool) : FullnessState := + let s : UInt8 := s + (s &&& (0b11111110 : UInt8)) ||| isFullAfter.toUInt8 + +/-- Whether resolving a document is guaranteed to fail in the given `FullnessState`. -/ +abbrev FailureCond := FullnessState → Bool + +/-- Input document consumed by the formatter, which chooses an optimal rendering of the document. -/ +inductive Doc where + /-- + Indicates that rendering this document is impossible. The formatter will always choose a rendering + of the document without `failure` nodes if one is available. + This is sometimes useful when defining custom combinators on a pre-existing document. + + Used when a `flattened` document contains a `newline none`. + + Example: + ``` + either (text "a") failure + ``` + produces + ``` + a + ``` + -/ + | failure + /-- + Designates a newline in the document. + + Within `flattened`, all `newline (some f)` nodes are replaced with `text f` + and all `newline none` nodes are replaced with `failure`, i.e. `newline none` can never be + flattened. + + Any newline that is not flattened by an outer `flattened` node will yield `\n` followed by + an amount of spaces corresponding to the current level of indentation as set by + `indent`, `align` and `reset` in the rendered document. + + `f?` is irrelevant during formatting: before formatting, a preprocessing step eliminates + all `flattened` nodes by replacing all `newline f?` nodes within each `flattened` node. + + Examples: + + ``` + indent 2 + (append + (append + (text "a") + (newline (some " "))) + (text "b")) + ``` + produces + ``` + a + b + ``` + --- + ``` + flattened + (append + (append + (text "a") + (newline (some " "))) + (text "b")) + ``` + produces + ``` + a b + ``` + --- + ``` + flattened + (append + (append + (text "a") + (newline none)) + (text "b")) + ``` + produces + ``` + + ``` + -/ + | newline (f? : Option String) + /-- + Designates a piece of text without newlines in the document. + `text` nodes are never broken apart by the formatter. + + The formatter will never choose a rendering where a `text` node is placed on the same line + after a `full` node. + + Examples: + + ``` + text "foo" + ``` + produces + ``` + foo + ``` + --- + ``` + either + (append + (full (text "a")) + (text "b")) + (text "c") + ``` + produces + ``` + c + ``` + -/ + | text (s : String) + /-- + Flattens an inner document by replacing all `newline (some f)` nodes in the inner + document with `text f` and all `newline none` nodes in the inner document with `failure`. + + `flattened` is eliminated before formatting by a preprocessing step that replaces all + `newline f?` nodes within each `flattened` node. + + Examples: + + ``` + flattened + (append + (append + (text "a") + (newline (some " "))) + (text "b")) + ``` + produces + ``` + a b + ``` + --- + ``` + flattened + (append + (append + (text "a") + (newline none)) + (text "b")) + ``` + produces + ``` + + ``` + -/ + | flattened (d : Doc) + /-- + Adds `n` to the current level of indentation within an inner document. + Multiple `indented` nodes with `isCumulative = false` will only increase the level of indentation + once per line, which is useful to request that an inner document needs to be indented if it + is broken over to another line, without these requests doubling up over multiple nested documents, + all of which may need to be indented once if they are all on separate lines, but should not be + indented several times if they end up on the same line. For non-cumulative `indented n ..` nodes + on the same line, the innermost `n` is used. + + When rendering a newline, the formatter produces an amount of spaces corresponding to the + current level of indentation + (i.e. the level of cumulative indentation + the level of non-cumulative indetation) + after the newline. + + Examples: + + ``` + indented 2 true + (append + (text "a") + (indented 2 true + (append + (append + (text "b") + (newline (some " "))) + (text "c")))) + ``` + produces + ``` + ab + c + ``` + --- + ``` + indented 2 false + (append + (text "a") + (append + (text "b") + (indented 2 false + (append + (newline (some " ")) + (text "c"))))) + ``` + produces + ``` + ab + c + ``` + --- + ``` + indented 2 false + (append + (text "a") + (append + (newline (some " ")) + (append + (text "b") + (indented 2 false + (append + (newline (some " ")) + (text "c")))))) + ``` + produces + ``` + a + b + c + ``` + -/ + -- Non-cumulative indentation does not exist in the Racket implementation, but it is very useful + -- when formatting Lean documents, as each line break in a nested term should only increase the + -- level of indentation by 1. + | indented (n : Nat) (isCumulative : Bool) (d : Doc) + /-- + Sets the current level of indentation within an inner document to the current column position + at the position where the `aligned` is rendered. + + When rendering a newline, the formatter produces an amount of spaces corresponding to the + current level of indentation after the newline. + + Example: + + ``` + append + (text "a") + (aligned + (append + (newline (some " ")) + (text "b"))) + ``` + produces + ``` + a + b + ``` + -/ + | aligned (d : Doc) + /-- + Sets the current level of indentation within an inner document to 0. + + When rendering a newline, the formatter produces an amount of spaces corresponding to the + current level of indentation after the newline. + + Example: + + ``` + indented 2 + (append + (unindented + (append + (text "a") + (newline (some " ")))) + (text "b")) + ``` + produces + ``` + a + b + ``` + -/ + | unindented (d : Doc) + /-- + Enforces that no text can be placed on the same line after the inner document. + + Example: + + ``` + either + (append + (full (text "a")) + (text "b")) + (text "c") + ``` + produces + ``` + c + ``` + -/ + | full (d : Doc) + /-- + Designates a document that can be rendered to one of two alternatives. + + The formatter will always choose a non-failing alternative if one is available or fail otherwise. + When both alternatives are not failing, it chooses an optimal rendering from both alternatives. + + If the two subtrees of an `either` have the same structure, then this structure should be + referentially shared between the two subtrees instead of duplicating them. This ensures that + documents with lots of alternatives can still be formatted efficiently, as the formatter will be + able to re-use state across these alternatives. + + Examples: + + ``` + either (text "a") failure + ``` + produces + ``` + a + ``` + --- + ``` + either + (text "a") + (append + (append + (text "b") + (newline (some " "))) + (text "c")) + ``` + (assuming a lawful cost function) produces + ``` + a + ``` + -/ + | either (a b : Doc) + /-- + Appends the second document to the last line of the first document. + + Example: + + ``` + append + (append + (append + (text "a") + (newline (some " "))) + (text "b")) + (text "c") + ``` + produces + ``` + a + bc + ``` + -/ + | append (a b : Doc) +with + /-- + Determines whether resolving the document is guaranteed to fail in the given `FullnessState`. + -/ + @[computed_field] isFailure : Doc → FailureCond + -- `failure` always fails. All resolutions that contain `failure` can be pruned. + | .failure => fun _ => true + -- `newline` starts a new line, which can never be full at this point. + -- Hence, resolutions in which `isFullAfter` is true directly after `newline` can be pruned. + | .newline .. => (·.isFullAfter) + | .text s => fun state => + match state.isFullBefore, state.isFullAfter with + -- `text` nodes can be placed on non-full lines. + | false, false => false + -- `text` nodes cannot turn a line from being full to non-full. + | true, false => true + -- `text` nodes cannot turn a line from being non-full to full. + | false, true => true + -- Empty text nodes can be inserted on a full line, while non-empty text nodes cannot. + | true, true => ! s.isEmpty + -- `full` designates that the line is full. + -- Hence, resolutions in which `isFullAfter` is false directly after `full` can be pruned. + | .full _ => (! ·.isFullAfter) + -- For all of the remaining inner nodes, whether resolving the document is guaranteed to fail + -- depends on the child nodes below the inner node. + | _ => fun _ => false + /-- + Designates an overapproximation for the amount of newlines in a document. + This is used by the formatter to choose renderings amongst multiple alternatives + that all exceed a maximum optimality cutoff width, which bounds the total search space. + -/ + @[computed_field] maxNewlineCount? : Doc → Option Nat + | .failure => none + | .newline .. => some 1 + | .text _ + | .flattened _ => some 0 + | .indented _ _ d + | .aligned d + | .unindented d + | .full d => maxNewlineCount? d + | .either a b => .merge (max · ·) (maxNewlineCount? a) (maxNewlineCount? b) + | .append a b => .merge (· + ·) (maxNewlineCount? a) (maxNewlineCount? b) + /-- + Memoization height of this document. Documents with a `memoHeight` of 0 are memoized when + formatting. + Time-complexity-wise, this is sound because the formatting document is a binary tree, and so its + height bounds the amount of nodes in the tree. + -/ + @[computed_field] memoHeight : Doc → Nat + | .failure + | .newline .. + | .text _ => memoHeightLimit + | .flattened d + | .indented _ _ d + | .aligned d + | .unindented d + | .full d => + let n := memoHeight d + nextMemoHeight n + | .either a b + | .append a b => + let n := min (memoHeight a) (memoHeight b) + nextMemoHeight n +deriving Inhabited, Repr + +/-- +Designates a document that either contains all newlines in an inner document or where all newlines +have been flattened. + +The formatter will always choose a non-failing alternative if one is available or fail otherwise. +When both alternatives are not failing, it chooses an optimal rendering from both alternatives. + +`maybeFlattened d` is equivalent to `either d (flattened d)`. + +This construct corresponds to `group` in most traditional formatting languages. +-/ +def Doc.maybeFlattened (d : Doc) : Doc := + .either d d.flattened + +/-- +Designates a newline that is flattened to a single space when placed inside of a `flattened` node. + +Equivalent to `newline (some " ")`. +-/ +def Doc.nl : Doc := + .newline (some " ") + +/-- +Designates a newline that is flattened to an empty string when placed inside of a `flattened` node. + +Equivalent to `newline (some "")`. +-/ +def Doc.break : Doc := + .newline (some "") + +/-- +Designates a newline that cannot be flattened and will produce a `failure` node when attempting +to flatten it. + +Equivalent to `newline none`. +-/ +def Doc.hardNl : Doc := + .newline none + +/-- +Ensures that the level of indentation of an inner document is increased by 2 spaces after a newline. +Multiple `nested` nodes on the same line only increase the level of indentation once. + +Examples: + +``` +nested + (append + (text "a") + (append + (text "b") + (nested + (append + nl + (text "c"))))) +``` +produces +``` +ab + c +``` +--- +``` +nested + (append + (text "a") + (append + nl + (append + (text "b") + (nested + (append + nl + (text "c")))))) +``` +produces +``` +a + b + c +``` +-/ +def Doc.nested (d : Doc) : Doc := + .indented 2 false d + +/-- +Increases the level of indentation of an inner document by 2 spaces after a newline. +Multiple `hardNested` nodes on the same line each increase the level of indentation by 2. + +Example: + +``` +hardNested + (append + (text "a") + (hardNested + (append + (append + (text "b") + nl) + (text "c")))) +``` +produces +``` +ab + c +``` +-/ +def Doc.hardNested (d : Doc) : Doc := + .indented 2 true d + +/-- +Designates a document that can be rendered to one of several alternatives. + +The formatter will always choose a non-failing alternative if one is available or fail otherwise. +When more than one alternative is not failing, it chooses an optimal rendering from +the non-failing alternatives. + +Equivalent to `failure` if the set of alternatives is empty. +-/ +def Doc.oneOf (ds : Array Doc) : Doc := + match ds[0]? with + | none => + .failure + | some d => + ds[1:].foldl (init := d) fun acc d => acc.either d + +/-- +Appends multiple documents. Each document is appended to the last line of the preceding document. +-/ +def Doc.join (ds : Array Doc) : Doc := + match ds[0]? with + | none => + .text "" + | some d => + ds[1:].foldl (init := d) fun acc d => acc.append d + +/-- +Appends multiple documents with a separator document between each pair of adjacent documents. +-/ +def Doc.joinUsing (sep : Doc) (ds : Array Doc) : Doc := + match ds[0]? with + | none => + .text "" + | some d => + ds[1:].foldl (init := d) fun acc d => acc.append sep |>.append d + +/-- +Appends multiple documents with a separator document between each pair of adjacent documents with +optional newlines between them. +-/ +def Doc.fillUsing (sep : Doc) (ds : Array Doc) : Doc := Id.run do + let some hd := ds[0]? + | return .text "" + let mut r : Doc := hd + for d in ds do + r := Doc.either + (Doc.join #[r, sep, d]) + (Doc.join #[r, sep, .hardNl, d]) + return r + +instance : Append Doc where + append d1 d2 := d1.append d2 + +class ToDoc (α : Type) where + toDoc : α → Doc + +instance : ToDoc Doc where + toDoc d := d + +instance : ToDoc String where + toDoc s := .text s + +syntax:max "fmt!" interpolatedStr(term) : term + +macro_rules + | `(fmt! $interpStr) => do interpStr.expandInterpolatedStr (← `(Doc)) (← `(ToDoc.toDoc)) diff --git a/src/Lean/Data/Fmt/Formatter.lean b/src/Lean/Data/Fmt/Formatter.lean new file mode 100644 index 000000000000..2470574454b2 --- /dev/null +++ b/src/Lean/Data/Fmt/Formatter.lean @@ -0,0 +1,1012 @@ +/- +Copyright (c) 2025 Lean FRO. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. + +Author: Marc Huisinga +-/ +module + +prelude +public import Lean.Data.Fmt.Basic +public import Std.Data.HashSet.Basic + +/-! +`Fmt` formatter. + +This file implements the formatter of 'A Pretty Expressive Printer' [1] by +Sorawee Porncharoenwase, Justin Pombrio and Emina Torlak. +This implementation is based on the Racket implementation of pretty-expressive [2]. + +[1] https://arxiv.org/pdf/2310.01530 +[2] https://docs.racket-lang.org/pretty-expressive/ +-/ + +namespace Lean.Fmt + +open Std + +/-- +Whether the formatter should memoize the given document. +Since memoization in itself can be expensive, not all documents are memoized, only every n-th layer. +Time-complexity-wise, this is sound because the formatting document is a binary tree, and so its +height bounds the amount of nodes in the tree. +-/ +def Doc.shouldMemoize (d : Doc) : Bool := + d.memoHeight = 0 + +structure PreprocessingCacheKey where + docPtr : USize + isFlattened : Bool + deriving BEq, Hashable + +structure PreprocessingState where + cache : HashMap PreprocessingCacheKey Doc := {} + +/-- +Erases all `flattened` nodes from a document by flattening all newlines with each `flattened` node. + +The important property we require of `preprocess` is that it does not destroy the sharing in the +input document: a document of shared size n must still be of shared size O(n) after preprocessing. +This ensures that preprocessed documents can still be formatted asymptotically as quickly as the +input document. + +Notably, preprocessing `flattened` nodes does not destroy the sharing of the input document, as +each document occurs at most in its flattened or non-flattened form, and so each document is +duplicated at most once. + +Eliminating `indented`, `aligned` and `unaligned` nodes by computing the indentation level of each +leaf node and then reducing `newline` nodes to an unindented `newline` node and some text +representing the current level of indentation is not possible for this reason, +as each document can occur in arbitrarily many indentation contexts, and so the sharing of the +input document would be destroyed. + +The Racket implementation skips this step by implementing a global preprocessing cache +and implementing `flattened` as a function that flattens the newlines in the inner document. +We instead implement this as a separate preprocessing step to circumvent the global +preprocessing cache. +-/ +def Doc.preprocess (d : Doc) : Doc := + goMemoized d false |>.run' {} +where + goMemoized (d : Doc) (isFlattened : Bool) : StateM PreprocessingState Doc := do + let cacheKey := { docPtr := unsafe ptrAddrUnsafe d, isFlattened } + -- Re-using cached preprocessing results is essential for not destroying the + -- shared structure of the input document. + if let some d' := (← get).cache.get? cacheKey then + return d' + let d' ← go d isFlattened + modify fun s => { s with cache := s.cache.insert cacheKey d' } + return d' + go (d : Doc) (isFlattened : Bool) : StateM PreprocessingState Doc := do + match d with + | .newline f? => + if isFlattened then + let some f := f? + | return .failure + return .text f + else + return .newline none + | .flattened d => + goMemoized d true + | .failure + | .text .. => + return d + | .indented n c d => + let d ← goMemoized d isFlattened + return .indented n c d + | .aligned d => + let d ← goMemoized d isFlattened + return .aligned d + | .unindented d => + let d ← goMemoized d isFlattened + return .unindented d + | .full d => + let d ← goMemoized d isFlattened + return .full d + | .either d1 d2 => + let d1 ← goMemoized d1 isFlattened + let d2 ← goMemoized d2 isFlattened + return .either d1 d2 + | .append d1 d2 => + let d1 ← goMemoized d1 isFlattened + let d2 ← goMemoized d2 isFlattened + return .append d1 d2 + +/-- +Cost function that the formatter is invoked with. + +Must satisfy the laws documented in `LawfulCost`. +-/ +public class Cost (τ : Type) [Add τ] [LE τ] where + /-- Cost of inserting a text of `length` at `columnPos`. -/ + textCost : (columnPos length : Nat) → τ + /-- Cost of inserting a newline with `indentationAfterNewline`. -/ + newlineCost : (indentationAfterNewline : Nat) → τ + /-- + Maximum width after which the formatter stops trying to find an optimal rendering + according to the cost function and instead reverts to simpler heuristics to choose a rendering. + This value should be chosen to be larger than the actual column limit so that the formatter + can produce optimal renderings even when all renderings exceed the column limit. + -/ + optimalityCutoffWidth : Nat + +/-- +A measure is a tuple of the compound cost of a specific rendering and a writer monad to produce the +rendering. + +The compound cost of a measure consists of both a configurable cost (as defined by a configurable +cost function) and the current length of the last line of the rendering. +A measure is considered to be cheaper than (or to dominate) another measure if both the configurable +cost as determined by the cost function and the last length of the line are smaller than that of +the other measure. In the formatter, we prune measures if they are dominated by another measure. + +For a lawful cost function, the configurable cost of a measure increases with more lines and +as lines get longer, i.e. it increases as documents are appended to it. +This means that we cannot simply prune measures using the configurable cost alone: +a measure might have a lower configurable cost than another measure for the time being, but when we +append to both measures, the second measure might suddenly become more expensive than the first one. + +With the default cost function, this occurs if e.g. both renderings have the same amount of lines, +all of which are below the column limit, while the second rendering is close to the column limit on +the last line. Appending lots of text to the last line of both renderings will then cause the cost +of the second measure to balloon relative to that of the first one. + +Notably, this kind of future divergence in cost between the two measures is limited to the last line +of the rendering, as we will always append the exact same documents to both of them and their column +positions will be synced when a newline is inserted. Additionally, lawful cost functions have the +property that inserting text at a smaller column position (i.e. at the end of a shorter last line) +will always be cheaper than inserting text at a larger column position, and so a compound cost that +is smaller both w.r.t. the configurable cost and the last line length than another compound cost +will also remain smaller than the other cost in the future when we append documents to the last +line, which means that we can safely prune the dominated measure. + +In summary, for a lawful cost function, it is both necessary and sufficient to include the length of +the last line as a separate parameter in the compound cost and only prune measures that dominate +other measures: +- It is necessary because not including it can cause us to prune measures that will become cheaper + than other measures in the future +- It is sufficient because the future divergence in cost for a lawful cost function is limited to + the last line of the rendering, and for lawful cost functions inserting text at a smaller + column position (i.e. at the end of a shorter last line) will always be cheaper than inserting + at a larger column position. + +Finally, the inclusion of the last line length in the compound cost bounds the time complexity of +the formatter by bounding the maximum size of the sets of measures it processes: +Each cost function comes with an optimality cutoff width `W`, after which the formatter will stop +attempting to compute optimal measures according to the configurable cost and simply pick just one +heuristically. Hence, all measures in a set of measures that do not exceed `W` have a +last line length of at most `W`. +When sets of measures are combined by the formatter, it prunes dominated measures to retain the +invariant that sets of measures contain no dominated measures. +Together, this means that each set of measures in the formatter can only contain at most `W` +measures that don't dominate one another: if there were more than `W` measures, at least two +measures `m₁` and `m₂` must have the same last line length, which, by the totality of `≤` of +a lawful cost function, means that either `m₁` dominates `m₂`, or `m₂` dominates `m₁`. +-/ +structure Measure (τ : Type) where + /-- Length of the last line of the rendering represented by this measure. -/ + lastLineLength : Nat + /-- + Configurable cost (as definited by the cost function) of the rendering represented by this + measure. + -/ + cost : τ + /-- + Level of indentation after the rendering represented by this measure. + Non-cumulative indentation is persisted after a newline by increasing the level of indentation. + -/ + indentation : Nat + /-- + Level of non-cumulative indentation after the rendering represented by this measure. + Set to 0 after a newline. + -/ + nonCumulativeIndentation : Nat + /-- + Writer monad that produces the rendering that this measure presents. + -/ + output : StateM String Unit + +variable {τ : Type} [Add τ] [LE τ] [DecidableLE τ] [Cost τ] + +/-- +Whether a measure subsumes another measure. See the documentation of `Measure` for details on +what measure domination entails. +-/ +def Measure.dominates (m1 m2 : Measure τ) : Bool := + m1.lastLineLength <= m2.lastLineLength && m1.cost <= m2.cost + +/-- Determines the measure that represents the concatenation of the renderings of two measures. -/ +def Measure.append (m1 m2 : Measure τ) : Measure τ where + lastLineLength := m2.lastLineLength + cost := m1.cost + m2.cost + indentation := m2.indentation + nonCumulativeIndentation := m2.nonCumulativeIndentation + output := do + m1.output + m2.output + +/-- Adjusts the indentation and non-cumulative indentation of a measure. -/ +def Measure.adjustIndentation (m : Measure τ) (newIndentation : Nat) + (newNonCumulativeIndentation : Nat) : Measure τ := { m with + indentation := newIndentation + nonCumulativeIndentation := newNonCumulativeIndentation +} + + +/-- Runs the writer monad of a measure, printing its rendering to a string. -/ +def Measure.print (m : Measure τ) : String := + let (_, printed) := m.output.run "" + printed + +/-- +A tainted measure is a measure for a rendering that exceeds the optimality cutoff width of the +cost function passed to the formatter. + +Notably, it does not possess a compound cost that we maintain, but merely a series of steps that +describe how to resolve the tainted measure to a single measure, as well as an approximation of the +amount of newlines in the rendering of the tainted measure. + +The formatter will always prune tainted measures in favor of non-tainted measures. +When the formatter has to choose amongst multiple tainted measures, instead of tracking all of them +using a cost function like for non-tainted measures, it simply picks the tainted measure with the +largest approximation for the amount of newlines, so as to attempt to heuristically produce +renderings that are higher (in terms of amount of lines) instead of ones where all text is +squished into the same line. + +Tainting measures instead of attempting to determine an optimal one amongst multiple tainted +measures bounds the time complexity of the formatter, as described in the documentation +of `Measure`. + +Compared to the Racket implementation of pretty-expressive, `TaintedMeasure` is a defunctionalized +implementation of the tainted measures in the Racket implementation, which implements them using +promises that lazily resolve a tainted measure to a regular measure after the measure resolution +loop is complete. This implementation using promises violates the positivity constraints of +inductive types, as the lazy measure resolution would itself maintain a memoization cache that +can contain tainted measures. Defunctionalization ensures that we limit the set of potential +lazy resolutions to a finite set of (sound) options, which makes the type satisfy the positivity +constraint. +-/ +inductive TaintedMeasure (τ : Type) where + /-- + Merge two tainted measures. Resolving this tainted measure amounts to resolving the first measure + and only resolving the second measure if the resolution of the first tainted measure failed. + + Since there are only four different fullness states in which each document can be resolved and + potentially fail, since the failure of resolution is independent of column position and + indentation, and since the resolver for tainted measures memoizes whether a resolution failed, + the resolver for tainted measures will only need to try resolving at most `4*amount of documents` + alternatives overall, so the time complexity of the formatter remains bounded. + -/ + | mergeTainted (tm1 tm2 : TaintedMeasure τ) (maxNewlineCount? : Option Nat) + /-- + Append a document to the rendering of a tainted measure. Resolving this tainted measure amounts to + resolving the tainted measure on the left, resolving the document on the right in the column + position after resolving the tainted measure on the left and with the given + fullness state, picking a measure from the set of measures of the resolution on the right + and then appending those. + -/ + | taintedAppend (tm1 : TaintedMeasure τ) (d2 : Doc) (fullness : FullnessState) + (maxNewlineCount? : Option Nat) + /-- + Append a tainted measure to a regular measure. Resolving this tainted measure amounts to simply + resolving the tainted measure on the right and appending it to the measure on the left. + -/ + | appendTainted (m1 : Measure τ) (tm2 : TaintedMeasure τ) (maxNewlineCount? : Option Nat) + /-- + Change the level of indentation of a tainted measure. Resolving this tainted measure amounts to + resolving the tainted measure and adjusting the resulting indentation levels. + -/ + | adjustTaintedIndentation (tm : TaintedMeasure τ) + (newIndentation newNonCumulativeIndentation : Nat) (maxNewlineCount? : Option Nat) + /-- + Resolve a tainted measure for a given resolution context to a regular measure. + Amounts to resolving the given document in the given context, picking a measure from the set of + measures produced by the resolution and memoizing whether the resolution failed so that + no failed resolution of a tainted measure is tried twice in the same fullness state and the time + complexity for tainted measure resolution remains bounded by `4*amount of documents`. + + Notably, the resolution of the document in the given context skips the taintedness-check for the + top level node, so this will process the top-level node of the document and then recurse with + potentially tainted child documents until eventually the full tainted measure is resolved. + -/ + | resolveTainted (d : Doc) (columnPos : Nat) (indentation nonCumulativeIndentation : Nat) + (fullness : FullnessState) (maxNewlineCount? : Option Nat) + +/-- Approximation for the maximum amount of newlines in the rendering of a tainted measure. -/ +def TaintedMeasure.maxNewlineCount? : TaintedMeasure τ → Option Nat + | .mergeTainted (maxNewlineCount? := n) .. => n + | .taintedAppend (maxNewlineCount? := n) .. => n + | .appendTainted (maxNewlineCount? := n) .. => n + | .adjustTaintedIndentation (maxNewlineCount? := n) .. => n + | .resolveTainted (maxNewlineCount? := n) .. => n + +/-- +Yields a `TaintedMeasure.mergeTainted` where the tainted measure with a larger newline approximation +is resolved first. + +Yields just the measure with a larger newline approximation if `prunable` is set to `true`, which +should only be set if it can be guaranteed that both tainted measures will always fail at the same +time (in which case we never need to try both). +-/ +def TaintedMeasure.merge (tm1 tm2 : TaintedMeasure τ) (prunable : Bool) : TaintedMeasure τ := + let (tm1, tm2) := + if Option.le (· <= ·) tm2.maxNewlineCount? tm1.maxNewlineCount? then + (tm1, tm2) + else + (tm2, tm1) + if prunable then + tm1 + else + -- There are two reasonable options for this newline approximation: + -- 1. The newline approximation of the first measure (as used by the Racket implementation) + -- 2. The maximum of both newline approximations + -- The first option is more accurate if resolving `tm1` does not fail, in which case the second + -- option is a worse approximation, while the second option is more accurate if resolving + -- `tm1` can fail. + .mergeTainted tm1 tm2 tm1.maxNewlineCount? + +/-- +Set of non-tainted measures. + +Fulfills the following invariants: +1. No two measures dominate each other. +2. The set of non-tainted measures is sorted by last line length (strictly descending). + +Together, these two invariants also imply that the set of non-tainted measures is sorted +by cost (strictly ascending), as otherwise the first invariant would be violated. + +Since all of these measures are non-tainted, both invariants individually imply that there are +at most W measures in a given set of non-tainted measures, where W is the optimality cutoff width +of the cost function. +-/ +abbrev MeasureSet.Set (τ : Type) := List (Measure τ) + +/-- Merges two sets of non-tainted measures, maintaining their invariants in the result. -/ +partial def MeasureSet.Set.merge (ms1 ms2 : MeasureSet.Set τ) : MeasureSet.Set τ := + match ms1, ms2 with + | [], _ => ms2 + | _, [] => ms1 + | m1 :: ms1', m2 :: ms2' => + if m1.dominates m2 then + merge ms1 ms2' + else if m2.dominates m1 then + merge ms1' ms2 + else if m1.lastLineLength > m2.lastLineLength then + m1 :: merge ms1' ms2 + else + m2 :: merge ms1 ms2' + +/-- +A set of measures is either a single tainted measure or a set of non-tainted measures. +The formatter prefers non-empty sets of measures over tainted measures and tainted measures +over empty sets of measures. +-/ +inductive MeasureSet (τ : Type) where + | tainted (tm : TaintedMeasure τ) + | set (ms : MeasureSet.Set τ) + deriving Inhabited + +/-- +Merges two sets of measures, preferring non-empty sets of measures over tainted measures and tainted +measures over empty sets of measures. +Tainted measures are merged according to `TaintedMeasure.merge` and sets of non-tainted measures +are merged according to `MeasureSet.Set.merge`. + +`prunable` can only be set to `true` if either `ms1` and `ms2` are not both tainted, or if it can be +guaranteed that both tainted measures will always fail at the same time +(in which case we never need to try both). +-/ +def MeasureSet.merge (ms1 ms2 : MeasureSet τ) (prunable : Bool) : MeasureSet τ := + match ms1, ms2 with + | _, .set [] => + ms1 + | .set [], _ => + ms2 + | .tainted tm1, .tainted tm2 => + .tainted (tm1.merge tm2 prunable) + | _, .tainted _ => + ms1 + | .tainted _, _ => + ms2 + | .set ms1, .set ms2 => + .set (ms1.merge ms2) + +/-- +Adjusts all indentations in a set of measures according to `Measure.adjustIndentation` and +`TaintedMeasure.adjustTaintedIndentation`. +-/ +def MeasureSet.adjustIndentation (m : MeasureSet τ) (newIndentation : Nat) + (newNonCumulativeIndentation : Nat) : MeasureSet τ := + match m with + | .set ms => + .set <| ms.map (·.adjustIndentation newIndentation newNonCumulativeIndentation) + | .tainted tm => .tainted + (.adjustTaintedIndentation tm newIndentation newNonCumulativeIndentation tm.maxNewlineCount?) + +/-- +Memoization key for sets of measures produced by the formatter. +Includes the full context that uniquely determines a set of measures: +- A pointer to the document that is being formatted +- The column position at which the document is being formatted +- The current level of indentation within which the document is being formatted +- The current level of non-cumulative indentation within which the document is being formatted +- The fullness state surrounding the document +-/ +structure SetCacheKey where + docPtr : USize + columnPos : Nat + indentation : Nat + nonCumulativeIndentation : Nat + fullness : FullnessState + deriving BEq, Hashable + +/-- +Memoization key for tracking whether a document has failed in the resolver for tainted measures. +Since resolution failure only depends on the document and the fullness state surrounding it, +this key does not contain the column position or the current indentation level. + +Memoizing the failure state in the resolver for tainted measures ensures that we never have to +resolve a single document (as identified by its pointer) more than 4 times. +-/ +structure FailureCacheKey where + docPtr : USize + fullness : FullnessState + deriving BEq, Hashable + +/-- +State of the resolver and the resolver for tainted measures, which runs after the regular resolver. + +Maintains three separate memoization caches: +- `setCache`, which memoizes sets of measures that are produced during resolution per `SetCacheKey`. + This is the main memoization cache of the formatter. It memoizes all resolution results for all + documents with `shouldMemoize = true` and ensures that the time complexity of resolution remains + reasonable. + After resolution, the `setCache` is re-used in resolutions performed by the resolution of tainted + measures. Notably, in the resolution of tainted measures, it is not used for resolving the + top-level measure in a `TaintedMeasure.resolveTainted`, as this would simply again yield a + tainted measure, and no progress in resolving the tainted measure would be made. + In the Racket implementation, this cache is replaced by several mutable caches + (one per fullness state) on the document. +- `resolvedTaintedCache`, which memoizes the measure (if any) produced by resolving a tainted + measure. Tainted measures can be shared during resolution if they are cached in `setCache` and + then later re-used. This cache ensures that the resolver for tainted measures does not perform + additional work relative to the resolver if the resolver has already figured out that two tainted + measures are identical. + In the Racket implementation, this cache is replaced with mutable state on the tainted measure. +- `failureCache`, which memoizes whether resolving a document in a given fullness state resulted + in a failure. Resolution failure depends only on the document and the given fullness state that + the document is resolved in, so this cache allows pruning subtrees of the search more + aggressively. + In the resolver for tainted measures, this cache also ensures that we never try to resolve the + same document more than four times, which bounds the time complexity of the tainted resolver. + In the Racket implementation, this cache is a mutable cache on the document that is only used + in the resolver for tainted measures to bound its time complexity. However, we've found that + performance improves when also enabling it for the regular resolver. +-/ +structure ResolutionState (τ : Type) where + setCache : HashMap SetCacheKey (MeasureSet τ) := {} + resolvedTaintedCache : HashMap USize (Option (Measure τ)) := {} + failureCache : HashSet FailureCacheKey := {} + +/-- +Monad for resolving a document in a resolution context to a set of measures. +Uses `StateRefT` to avoid having to box the state together with return values during resolution. +-/ +abbrev ResolverM (σ τ : Type) := StateRefT (ResolutionState τ) (ST σ) + +def ResolverM.run (f : ResolverM σ τ α) : ST σ α := + f.run' {} + +@[inline] +def getCachedSet? (d : Doc) (columnPos indentation nonCumulativeIndentation : Nat) + (fullness : FullnessState) : + ResolverM σ τ (Option (MeasureSet τ)) := do + return (← get).setCache.get? { + docPtr := unsafe ptrAddrUnsafe d + columnPos + indentation + nonCumulativeIndentation + fullness + } + +@[inline] +def setCachedSet (d : Doc) (columnPos indentation nonCumulativeIndentation : Nat) + (fullness : FullnessState) (set : MeasureSet τ) : ResolverM σ τ Unit := + modify fun state => { state with + setCache := state.setCache.insert { + docPtr := unsafe ptrAddrUnsafe d + columnPos + indentation + nonCumulativeIndentation + fullness + } set + } + +inductive CacheResult (α : Type) where + | miss + | hit (cached : α) + +@[inline] +def getCachedResolvedTainted? (tm : TaintedMeasure τ) : + ResolverM σ τ (CacheResult (Option (Measure τ))) := do + match (← get).resolvedTaintedCache.get? (unsafe ptrAddrUnsafe tm) with + | none => return .miss + | some cached? => return .hit cached? + +@[inline] +def setCachedResolvedTainted (tm : TaintedMeasure τ) (m? : Option (Measure τ)) : + ResolverM σ τ Unit := + modify fun state => { state with + resolvedTaintedCache := state.resolvedTaintedCache.insert (unsafe ptrAddrUnsafe tm) m? + } + +def Doc.isLeaf : Doc → Bool + | .failure => true + | .newline .. => true + | .text .. => true + | _ => false + +def isFailing (d : Doc) (fullness : FullnessState) : ResolverM σ τ Bool := do + if d.isLeaf then + -- For leaf nodes, guaranteed failure is fully determinined by `Doc.isFailure`. + return d.isFailure fullness + else if d.isFailure fullness then + -- For some inner nodes (`full` specifically), we can prune specific subtrees + -- if `Doc.isFailure` yields `true` and have no information about failure otherwise. + return true + else + -- For all other nodes, if we have already determined that a document fails in a given fullness + -- state, we can prune that subtree. + let isCachedFailure := (← get).failureCache.contains { + docPtr := unsafe ptrAddrUnsafe d + fullness + } + return isCachedFailure + +def setCachedFailing (d : Doc) (fullness : FullnessState) : ResolverM σ τ Unit := + modify fun state => { state with + failureCache := state.failureCache.insert { + docPtr := unsafe ptrAddrUnsafe d + fullness + } + } + +def Resolver (σ τ : Type) := + (d : Doc) → (columnPos indentation nonCumulativeIndentation : Nat) → (fullness : FullnessState) → + ResolverM σ τ (MeasureSet τ) + +/-- +Checks whether we have a memoized result for a given resolution context and if so, uses that. +Otherwise, `f` is evaluated and the result is memoized if `Doc.shouldMemoize` is true. +-/ +@[specialize] +def Resolver.memoize (f : Resolver σ τ) : Resolver σ τ := + fun d columnPos indentation nonCumulativeIndentation fullness => do + if ← isFailing d fullness then + return .set [] + if columnPos > Cost.optimalityCutoffWidth τ || indentation > Cost.optimalityCutoffWidth τ + || ! d.shouldMemoize then + let r ← f d columnPos indentation nonCumulativeIndentation fullness + if r matches .set [] then + setCachedFailing d fullness + return r + if let some cachedSet ← + getCachedSet? d columnPos indentation nonCumulativeIndentation fullness then + return cachedSet + let r ← f d columnPos indentation nonCumulativeIndentation fullness + setCachedSet d columnPos indentation nonCumulativeIndentation fullness r + if r matches .set [] then + setCachedFailing d fullness + return r + +mutual + +/-- +Determines the set of measures for a given resolution context. +The root node is not memoized, while nodes below the root node can be memoized. + +Notably, this function skips checks that determine whether the context at the root node already +exceeds the optimality cutoff width, which (together with not memoizing the root node) means that +we can use this function to resolve tainted documents to non-tainted ones in the resolver for +tainted measures. +-/ +partial def MeasureSet.resolveCore : Resolver σ τ := + fun d columnPos indentation nonCumulativeIndentation fullness => do + match d with + | .failure => + return .set [] + | .newline .. => + let lineIndentation := indentation + nonCumulativeIndentation + return .set [{ + lastLineLength := lineIndentation + cost := Cost.newlineCost lineIndentation + -- Persist the non-cumulative indentation so that the following newlines use it. + indentation := indentation + nonCumulativeIndentation + -- Reset the level of non-cumulative indentation so that the next non-cumulative `indented` + -- can increase the level of indentation again. + nonCumulativeIndentation := 0 + output := modify fun out => + out ++ "\n" |>.pushn ' ' lineIndentation + }] + | .text s => + return .set [{ + lastLineLength := columnPos + s.length + cost := Cost.textCost columnPos s.length + indentation + nonCumulativeIndentation + output := modify fun out => + out ++ s + }] + | .flattened _ => + -- Eliminated during pre-processing. + unreachable! + | .indented n isCumulative d => + let ms ← + if isCumulative then + -- Increases the over-all level of indentation, which (if present) is spent from + -- the non-cumulative level of indentation. + -- This ensures that an outer non-cumulative `indented` will not duplicate the indentation + -- of an inner cumulative `intended`. + resolve d columnPos (indentation + n) (nonCumulativeIndentation - n) fullness + else + -- Sets the level of non-cumulative indentation to `n`. + -- In a chain of nested non-cumulative `indent`s, the innermost `n` is used. + resolve d columnPos indentation n fullness + return ms.adjustIndentation indentation nonCumulativeIndentation + | .aligned d => + -- Sets the level of indentation to `columnPos` and resets the level of + -- non-cumulative indentation, as the alignment dictates the level of indentation in `d`. + let ms ← resolve d columnPos columnPos 0 fullness + return ms.adjustIndentation indentation nonCumulativeIndentation + | .unindented d => + let ms ← resolve d columnPos 0 0 fullness + return ms.adjustIndentation indentation nonCumulativeIndentation + | .full d => + -- The failure condition of `full` ensures that `fullness.isFullAfter` is true when we reach + -- this point. However, within `full`, the `full` node imposes no constraints, so we case-split + -- on `fullness.isFullAfter` here. + let set1 ← resolve d columnPos indentation nonCumulativeIndentation + (fullness.setFullAfter false) + let set2 ← resolve d columnPos indentation nonCumulativeIndentation + (fullness.setFullAfter true) + return .merge set1 set2 (prunable := false) + | .either d1 d2 => + let set1 ← resolve d1 columnPos indentation nonCumulativeIndentation fullness + let set2 ← resolve d2 columnPos indentation nonCumulativeIndentation fullness + return .merge set1 set2 (prunable := false) + | .append d1 d2 => + -- We can't tell whether the line at the end of `d1` will be full in advance, which decides + -- whether we need to set `isFullAfter` on the left side of the `append` and `isFullBefore` + -- on the right side of the `append`, so we case-split on these two alternatives and then + -- later prune subtrees that are inconsistent with the given fullness state. + let set1 ← analyzeAppend d d1 d2 columnPos indentation nonCumulativeIndentation fullness false + let set2 ← analyzeAppend d d1 d2 columnPos indentation nonCumulativeIndentation fullness true + return .merge set1 set2 (prunable := false) +where + /-- + Resolves `d1` to a measure set, then resolves `d2` with each of the column positions in the + measure set of `d1` and finally appends each measure from resolving `d2` to each measure from + resolving `d1`. + At the end, the invariants for sets of measures (documented at `MeasureSet.Set`) are enforced. + -/ + analyzeAppend (d d1 d2 : Doc) (columnPos indentation nonCumulativeIndentation : Nat) + (fullness : FullnessState) (isMidFull : Bool) : ResolverM σ τ (MeasureSet τ) := do + let fullness1 := fullness.setFullAfter isMidFull + let fullness2 := fullness.setFullBefore isMidFull + let set1 ← resolve d1 columnPos indentation nonCumulativeIndentation fullness1 + match set1 with + | .tainted tm1 => + return .tainted (.taintedAppend tm1 d2 fullness2 d.maxNewlineCount?) + | .set ms1 => + ms1.foldrM (init := MeasureSet.set []) fun m1 acc => do + let set2 ← resolve d2 m1.lastLineLength m1.indentation m1.nonCumulativeIndentation fullness2 + let m1Result : MeasureSet τ := + match set2 with + | .tainted tm2 => + .tainted (.appendTainted m1 tm2 d.maxNewlineCount?) + | .set [] => + .set [] + | .set (m2 :: ms2) => .set <| Id.run do + let mut last := m1.append m2 + let mut deduped := [] + for m2 in ms2 do + let current := m1.append m2 + -- `ms2` fulfills the measure set invariants, which means that appending these + -- measures to `m1` results in a set that is still sorted by last line length + -- in strictly descending order. The resulting set is also still sorted by cost, but + -- in general not in strictly ascending order, just in ascending order + -- (by monotonicity of a lawful `+`): + -- A cost function over ℕ with `a + b := max(a, b)`, ≤ as on the natural numbers and + -- `textCost` / `newlineCost` defined in some arbitrary lawful manner is lawful, + -- while `ms2 := [(3, 1), (2, 2), (1, 3)]` fulfills the measure set invariants + -- and `map ms2 (append (2, 2) ·) = [(3, 2), (2, 2), (1, 3)]` is not strictly + -- ascending in cost. + -- The two order invariants imply that we only need to check adjacent measures for + -- domination and that we only need to check the cost of measures when checking for + -- domination. The fact that the cost order is not necessarily strict implies that + -- checking the cost of adjacent measures is still necessary for general lawful cost + -- functions. + if current.cost <= last.cost then + last := current + continue + deduped := last :: deduped + last := current + return last :: deduped |>.reverse + -- `m1Result` and (inductively) all results in `acc` are resolutions of `d2`, so all + -- resolutions being merged here either fail at once or none of them fail. + -- Hence, we can set `prunable := true` here. + return m1Result.merge acc (prunable := true) + +/-- +Determines the set of measures for a given resolution context and memoizes all nodes along the way. +-/ +partial def MeasureSet.resolve : Resolver σ τ := Resolver.memoize + fun d columnPos indentation nonCumulativeIndentation fullness => do + -- Lifting both the memoization of the root node and the taintedness check out to + -- `MeasureSet.resolve` ensures that we can use `resolveCore` to resolve `resolveTainted` nodes + -- in the resolver for tainted measures. + let columnPos' := + if let .text s := d then + columnPos + s.length + else + columnPos + if columnPos' > Cost.optimalityCutoffWidth τ || + indentation + nonCumulativeIndentation > Cost.optimalityCutoffWidth τ then + return .tainted + (.resolveTainted d columnPos indentation nonCumulativeIndentation fullness + d.maxNewlineCount?) + return ← resolveCore d columnPos indentation nonCumulativeIndentation fullness + +end + +def TaintedResolver (σ τ : Type) := + (tm : TaintedMeasure τ) → ResolverM σ τ (Option (Measure τ)) + +/-- +Checks whether we have a memoized result for a given tainted measure and if so, uses that. +Otherwise, `f` is evaluated and the result is memoized. + +We memoize all tainted resolution results because the resolver for tainted measures will only +have to resolve every document at most 4 times, as it only performs a case-split in `mergeTainted` +when one of the two resolutions fail, which is independent of indentation and column position and +only depends on the document and the fullness state surrounding it. +-/ +@[specialize] +def TaintedResolver.memoize (f : TaintedResolver σ τ) : TaintedResolver σ τ := fun tm => do + let cachedResolvedTainted? ← getCachedResolvedTainted? tm + if let .hit m := cachedResolvedTainted? then + return m + let m? ← f tm + setCachedResolvedTainted tm m? + return m? + +mutual + +partial def TaintedMeasure.resolve? : TaintedResolver σ τ := TaintedResolver.memoize + fun tm => do + match tm with + | .mergeTainted tm1 tm2 _ => + -- We need to try both alternatives here when the first alternative fails. + -- However, such failures only depend on the document and the surrounding fullness state, + -- so this will never try more than 4 separate alternatives per document overall, + -- which bounds the time complexity of the tainted resolver. + let some m1 ← tm1.resolve? + | let m2? ← tm2.resolve? + return m2? + return some m1 + | .taintedAppend tm d fullness _ => + let some m1 ← tm.resolve? + | return none + let ms2 ← MeasureSet.resolve d m1.lastLineLength m1.indentation m1.nonCumulativeIndentation fullness + let some m2 ← ms2.extractAtMostOne? + | return none + return some <| m1.append m2 + | .appendTainted m1 tm2 _ => + let some m2 ← tm2.resolve? + | return none + return some <| m1.append m2 + | .adjustTaintedIndentation tm newIndentation newNonCumulativeIndentation _ => + let some m ← tm.resolve? + | return none + return some <| m.adjustIndentation newIndentation newNonCumulativeIndentation + | .resolveTainted d columnPos indentation nonCumulativeIndentation fullness _ => + -- If we used `resolve` instead of `resolveCore` here, we would just again obtain a tainted + -- measure, and the mutual recursion between `MeasureSet.extractAtMostOne?` and + -- `TaintedMeasure.resolve?` would make no progress. + -- Using `resolveCore`, which does not perform taintedness checks and does not memoize the + -- result of resolving the root node, ensures that we make progress on the root node. + -- This resolution may again produce tainted measures for the children of `d`, which will then + -- be resolved recursively. + let ms ← MeasureSet.resolveCore d columnPos indentation nonCumulativeIndentation fullness + let m? ← ms.extractAtMostOne? + if m?.isNone then + setCachedFailing d fullness + return m? + +/-- +Yields the measure in a non-tainted measure set with the lowest cost and amongst measures with the +lowest cost, the one with the smallest last line length. +For a tainted measure, resolves the tainted measure to a regular measure. +-/ +partial def MeasureSet.extractAtMostOne? (ms : MeasureSet τ) : + ResolverM σ τ (Option (Measure τ)) := do + match ms with + | .tainted tm => + tm.resolve? + | .set ms => + return ms.head? + +end + +/-- +Resolves a document to a measure with the given initial offset, or `none` if the resolution +failed, i.e. if there is no interpretation of `d` that does not result in `failure`. +-/ +def resolve? (d : Doc) (offset : Nat) : Option (Measure τ) := runST fun _ => ResolverM.run do + -- We cannot tell in advance whether the last line of `d` will be full, so we case split on + -- `isFullAfter` of the fullness state and later prune subtrees of the search + -- when we notice that they are inconsistent with the actual document. + let ms1 ← MeasureSet.resolve d offset 0 0 (.mk false false) + let ms2 ← MeasureSet.resolve d offset 0 0 (.mk false true) + let ms := ms1.merge ms2 (prunable := false) + ms.extractAtMostOne? + +/-- +Formats a document to a string for a given cost function. +Yields `none` if the resolution failed, i.e. if there is no interpretation of `d` that does not +result in `failure`. +-/ +public def formatWithCost? (τ : Type) [Add τ] [LE τ] [DecidableLE τ] [Cost τ] + (d : Doc) (offset : Nat := 0) : Option String := do + let d := d.preprocess + let m ← resolve? (τ := τ) d offset + return m.print + +/-- +Default cost function for the formatter. + +Minimizes the sum of squared overflows over a page width limit `softWidth`. This means that the +formatter will find renderings with smaller overflows even when all possible renderings for a +document overflow the page width limit. +Amongst renderings with the same sum of squared overflows (or no overflows), it minimizes the +amount of newlines in the document. + +If the width of all renderings of a document exceed a parameter `optimalityCutoffWidth`, +the formatter will not attempt to determine an optimal rendering with the least amount of overflow +amongst these renderings. Instead, it heuristically chooses a rendering using an approximation for +the amount of newlines, and picks the rendering with the largest approximation for the amount of +newlines. + +`optimalityCutoffWidth` bounds the worst-case time complexity of the formatter. +It does not represent the actual page limit and should always be chosen to be larger than +`softWidth`. +-/ +public structure DefaultCost (softWidth : Nat) (optimalityCutoffWidth : Nat) where + widthCost : Nat + heightCost : Nat + +def DefaultCost.add (c1 c2 : DefaultCost w W) : DefaultCost w W := + ⟨c1.widthCost + c2.widthCost, c1.heightCost + c2.heightCost⟩ + +@[no_expose] +public instance : Add (DefaultCost w W) where + add := DefaultCost.add + +def DefaultCost.le + (c1 c2 : DefaultCost w W) : Bool := + if c1.widthCost = c2.widthCost then + c1.heightCost ≤ c2.heightCost + else + c1.widthCost ≤ c2.widthCost + +@[no_expose] +public instance : LE (DefaultCost w W) where + le c1 c2 := DefaultCost.le c1 c2 + +@[no_expose] +public instance : DecidableLE (DefaultCost w W) := + fun _ _ => inferInstanceAs (Decidable (_ = true)) + +def DefaultCost.textCost (softWidth optimalityCutoffWidth columnPos length : Nat) : + DefaultCost softWidth optimalityCutoffWidth := + if columnPos + length <= softWidth then + -- `softWidth` not exceeded => no cost + ⟨0, 0⟩ + else if columnPos <= softWidth then + -- `softWidth` first exceeded with this text node by `columnPos + length - softWidth` + -- => cost of `(columnPos + length - softWidth)^2` + let lengthOverflow := (columnPos + length) - softWidth + ⟨lengthOverflow*lengthOverflow, 0⟩ + else + -- This text node is being placed at a column position that already exceeds `softWidth`, + -- which means that we have already booked costs for another text node before this one on + -- the same line. + -- We want the sum of these two costs to represent the combined squared overflow over + -- `softWidth` so that the sum of all costs of the text nodes on a line denotes the total + -- squared overflow. + -- + -- Assume that the cost `c₁` of the text nodes that have already been placed on this line prior + -- to this one represents the squared overflow over `softWidth` so far, i.e. that + -- `c₁ = (columnPos - softWidth)^2` (the induction basis for this assumption is given by the + -- first two branches of this function). + -- + -- We now want to choose a cost `c₂` for this text node with + -- `c₁ + c₂ = (columnPos + length - softWidth)^2` to maintain the invariant. + -- With `columnPos > softWidth` and `(a + b)^2 - a^2 = b*(2*a + b)`, we have + -- ``` + -- c₁ + c₂ = (columnPos - softWidth)^2 + c₂ = (columnPos + length - softWidth)^2 iff + -- c₂ = (columnPos - softWidth + length)^2 - (columnPos - softWidth)^2 + -- = length*(2*(columnPos - softWidth) + length) + -- ```. + let columnPosOverflow := columnPos - softWidth + let lengthOverflow := length + ⟨lengthOverflow*(2*columnPosOverflow + lengthOverflow), 0⟩ + +def DefaultCost.newlineCost (w W _length : Nat) : + DefaultCost w W := + ⟨0, 1⟩ + +@[no_expose] +public instance : Cost (DefaultCost softWidth optimalityCutoffWidth) where + textCost := DefaultCost.textCost softWidth optimalityCutoffWidth + newlineCost := DefaultCost.newlineCost softWidth optimalityCutoffWidth + optimalityCutoffWidth := optimalityCutoffWidth + +/-- +Formats a document to a string with the default cost function for a given page width limit `width`. +Yields `none` if the resolution failed, i.e. if there is no interpretation of `d` that does not +result in `failure`. + +The default cost function minimizes the sum of squared overflows over `width`. This means that the +formatter will find renderings with smaller overflows even when all possible renderings for a +document overflow the page width limit. +Amongst renderings with the same sum of squared overflows (or no overflows), it minimizes the +amount of newlines in the document. + +If the width of all renderings of a document exceed `optimalityCutoffWidth`, +the formatter will not attempt to determine an optimal rendering with the least amount of overflow +amongst these renderings. Instead, it heuristically chooses a rendering using an approximation for +the amount of newlines, and picks the rendering with the largest approximation for the amount of +newlines. + +`optimalityCutoffWidth` bounds the worst-case time complexity of the formatter. +It does not represent the actual page limit and should always be chosen to be larger than +`width`. +-/ +public def format? (d : Doc) (width : Nat) + (optimalityCutoffWidth : Nat := Nat.max ((5*width)/4) 200) + (offset : Nat := 0) : + Option String := do + formatWithCost? (τ := DefaultCost width optimalityCutoffWidth) d offset + +section DefaultCostDefTheorems + +public theorem DefaultCost.add_def {c₁ c₂ : DefaultCost w W} : + c₁ + c₂ = ⟨c₁.widthCost + c₂.widthCost, c₁.heightCost + c₂.heightCost⟩ := by + simp only [HAdd.hAdd, Add.add, DefaultCost.add] + +public theorem DefaultCost.le_def {c₁ c₂ : DefaultCost w W} : + c₁ ≤ c₂ ↔ + (if c₁.widthCost = c₂.widthCost then + c₁.heightCost ≤ c₂.heightCost + else + c₁.widthCost ≤ c₂.widthCost) := by + simp only [LE.le] + simp only [le, Bool.ite_eq_true_distrib, decide_eq_true_eq, Nat.le_eq] + +public theorem DefaultCost.textCost_def : + (Cost.textCost cp n : DefaultCost w W) = + (if cp + n <= w then + ⟨0, 0⟩ + else if cp <= w then + let o := (cp + n) - w + ⟨o*o, 0⟩ + else + ⟨n*(2*(cp - w) + n), 0⟩) := by + simp only [Cost.textCost, textCost] + +public theorem DefaultCost.newlineCost_def : + (Cost.newlineCost n : DefaultCost w W) = ⟨0, 1⟩ := by + simp only [Cost.newlineCost, newlineCost] + +end DefaultCostDefTheorems diff --git a/src/Lean/Data/Fmt/Json.lean b/src/Lean/Data/Fmt/Json.lean new file mode 100644 index 000000000000..566f4c3fa60f --- /dev/null +++ b/src/Lean/Data/Fmt/Json.lean @@ -0,0 +1,81 @@ +/- +Copyright (c) 2025 Lean FRO. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. + +Author: Marc Huisinga +-/ +module + +prelude +public import Lean.Data.Fmt.Formatter +public import Lean.Data.Json + +public section + +namespace Lean.Fmt.Json + +def isPrimitive (j : Json) : Bool := + match j with + | .arr .. => false + | .obj .. => false + | _ => true + +partial def format (j : Json) : Doc := + match j with + | .null => .text "null" + | .bool b => .text (toString b) + | .num n => .text (toString n) + | .str s => .join #[.text "\"", .text s, .text "\""] + | .arr a => + let elems := a.map format + if a.all isPrimitive then + .join #[ + .text "[", + .aligned + (.fillUsing (.text ", ") elems), + .text "]" + ] + else + .join #[ + .text "[", + .hardNested + (.append + .hardNl + (.joinUsing + (.append (.text ",") .hardNl) + elems)), + .hardNl, + .text "]" + ] + | .obj kv => + let pairs := kv.toArray.map fun (k, v) => formatKvPair k v + .join #[ + .text "{", + .hardNested + (.append + .hardNl + (.joinUsing + (.append (.text ",") .hardNl) + pairs)), + .hardNl, + .text "}" + ] +where + formatKvPair (k : String) (v : Json) : Doc := + let v' := format v + let f1 := .join #[ + .text "\"", + .text k, + .text "\": ", + v' + ] + if isPrimitive v then + let f2 := .join #[ + .text "\"", + .text k, + .text "\":", + .hardNested (.append .hardNl v') + ] + .either f1 f2 + else + f1 diff --git a/src/Lean/Data/Fmt/LawfulCost.lean b/src/Lean/Data/Fmt/LawfulCost.lean new file mode 100644 index 000000000000..6661a89e6385 --- /dev/null +++ b/src/Lean/Data/Fmt/LawfulCost.lean @@ -0,0 +1,121 @@ +/- +Copyright (c) 2025 Lean FRO. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. + +Author: Marc Huisinga +-/ +module + +prelude +public import Lean.Data.Fmt.Formatter +public import Init.Grind.Module.Basic +import Init + +/-! +This file documents the properties that a cost function in the `Fmt` formatter must fulfill +in order for formatting with the cost function to be correct and efficient. +It also proves that the default cost function fulfills these properties. + +The properties in this file have been taken from 'A Pretty Expressive Printer' [1] by +Sorawee Porncharoenwase, Justin Pombrio and Emina Torlak. + +[1] https://arxiv.org/pdf/2310.01530 +-/ + +namespace Lean.Fmt + +/-- +`LawfulCost` documents the properties that a cost function in the `Fmt` formatter mus fulfill +in order for formatting with the cost function to be correct and efficient. +-/ +public class LawfulCost (τ : Type) [Add τ] [LE τ] extends Cost τ, Grind.AddCommMonoid τ, Std.IsLinearOrder τ where + zero := textCost 0 0 + + textCost_monotone (cp₁ cp₂ n : Nat) : + cp₁ ≤ cp₂ → textCost cp₁ n ≤ textCost cp₂ n + textCost_add (cp n₁ n₂ : Nat) : + textCost cp (n₁ + n₂) = textCost cp n₁ + textCost (cp + n₁) n₂ + newlineCost_monotone (i₁ i₂ : Nat) : + i₁ ≤ i₂ → newlineCost i₁ ≤ newlineCost i₂ + + add_monotone (c₁ c₂ c₃ c₄ : τ) : c₁ ≤ c₂ → c₃ ≤ c₄ → c₁ + c₃ ≤ c₂ + c₄ + +attribute [grind ext] DefaultCost + +def DefaultCost.zero : DefaultCost w W := + ⟨0, 0⟩ + +instance : Zero (DefaultCost w W) where + zero := DefaultCost.zero + +theorem DefaultCost.zero_def : (0 : DefaultCost w W) = ⟨0, 0⟩ := by + simp only [Zero.zero, OfNat.ofNat, DefaultCost.zero] + +theorem DefaultCost.add_zero (c : DefaultCost w W) : c + 0 = c := by + simp only [zero_def, add_def] + grind + +theorem DefaultCost.add_comm (c₁ c₂ : DefaultCost w W) : c₁ + c₂ = c₂ + c₁ := by + simp only [add_def] + grind + +theorem DefaultCost.add_assoc (c₁ c₂ c₃ : DefaultCost w W) : + (c₁ + c₂) + c₃ = c₁ + (c₂ + c₃) := by + simp only [add_def] + grind + +instance : Grind.AddCommMonoid (DefaultCost w W) where + zero := DefaultCost.zero + add_zero := DefaultCost.add_zero + add_comm := DefaultCost.add_comm + add_assoc := DefaultCost.add_assoc + +theorem DefaultCost.le_refl (c : DefaultCost w W) : c ≤ c := by + simp only [le_def] + grind + +theorem DefaultCost.le_trans (c₁ c₂ c₃ : DefaultCost w W) : c₁ ≤ c₂ → c₂ ≤ c₃ → c₁ ≤ c₃ := by + simp only [le_def] + grind + +theorem DefaultCost.le_antisymm (c₁ c₂ : DefaultCost w W) : c₁ ≤ c₂ → c₂ ≤ c₁ → c₁ = c₂ := by + simp only [le_def] + grind + +theorem DefaultCost.le_total (c₁ c₂ : DefaultCost w W) : c₁ ≤ c₂ ∨ c₂ ≤ c₁ := by + simp only [le_def] + grind + +instance : Std.IsLinearOrder (DefaultCost w W) where + le_refl := DefaultCost.le_refl + le_trans := DefaultCost.le_trans + le_antisymm := DefaultCost.le_antisymm + le_total := DefaultCost.le_total + +theorem DefaultCost.textCost_monotone + (cp₁ cp₂ n : Nat) : + cp₁ ≤ cp₂ → + (Cost.textCost cp₁ n : DefaultCost w W) ≤ + Cost.textCost cp₂ n := by + grind [= textCost_def, = le_def, Nat.mul_le_mul] + +theorem DefaultCost.textCost_add (cp n₁ n₂ : Nat) : + (Cost.textCost cp (n₁ + n₂) : DefaultCost w W) = + Cost.textCost cp n₁ + Cost.textCost (cp + n₁) n₂ := by + grind [= textCost_def, = add_def, Nat.sub_add_comm] + +theorem DefaultCost.newlineCost_monotone (i₁ i₂ : Nat) : + i₁ ≤ i₂ → + (Cost.newlineCost i₁ : DefaultCost w W) ≤ + Cost.newlineCost i₂ := by + grind [newlineCost_def] + +def DefaultCost.add_monotone (c₁ c₂ c₃ c₄ : DefaultCost w W) : c₁ ≤ c₂ → c₃ ≤ c₄ → c₁ + c₃ ≤ c₂ + c₄ := by + grind [= le_def, = add_def] + +instance : LawfulCost (DefaultCost softWidth optimalityCutoffWidth) where + textCost_monotone := DefaultCost.textCost_monotone + textCost_add := DefaultCost.textCost_add + newlineCost_monotone := DefaultCost.newlineCost_monotone + + add_monotone := DefaultCost.add_monotone diff --git a/tests/bench/fmtConcat.lean b/tests/bench/fmtConcat.lean new file mode 100644 index 000000000000..be1a374f6b75 --- /dev/null +++ b/tests/bench/fmtConcat.lean @@ -0,0 +1,29 @@ +import Lean.Data.Fmt.Formatter + +open Lean.Fmt + +def n := 1000 --0 + +def pp (n : Nat) : Doc := + if n = 0 then + .text "" + else + .append (pp (n - 1)) (.text "line") + +@[noinline] +def doc (n : Nat) : IO Doc := + return pp n + +@[noinline] +def format (doc : Doc) : IO (Option String) := do + return format? doc 80 100 + +def main (args : List String) : IO Unit := do + let n := (args[0]!).toNat! + let d ← doc n + let startNs ← IO.monoNanosNow + let r? ← format d + let endNs ← IO.monoNanosNow + let benchTime : Float := (endNs - startNs).toFloat / 1_000_000_000.0 + assert! r?.isSome + IO.println s!"format: {benchTime}" diff --git a/tests/bench/fmtFillSep.lean b/tests/bench/fmtFillSep.lean new file mode 100644 index 000000000000..5dbeecf5747c --- /dev/null +++ b/tests/bench/fmtFillSep.lean @@ -0,0 +1,33 @@ +import Lean.Data.Fmt.Formatter + +open Lean.Fmt + +def fillSep (xs : Array String) : Doc := Id.run do + let some hd := xs[0]? + | return .text "" + let mut r : Doc := .text hd + for x in xs do + r := Doc.either + (Doc.joinUsing (.text " ") #[r, .text x]) + (Doc.joinUsing .hardNl #[r, .text x]) + return r + +@[noinline] +def doc (n : Nat) : IO Doc := do + let words ← IO.FS.readFile "fmtFillSepWords" + let words := words.splitOn "\n" |>.take n |>.toArray + return fillSep words + +@[noinline] +def format (doc : Doc) : IO (Option String) := do + return format? doc 80 100 + +def main (args : List String) : IO Unit := do + let n := (args[0]!).toNat! + let d ← doc n + let startNs ← IO.monoNanosNow + let r? ← format d + let endNs ← IO.monoNanosNow + let benchTime : Float := (endNs - startNs).toFloat / 1_000_000_000.0 + assert! r?.isSome + IO.println s!"format: {benchTime}" diff --git a/tests/bench/fmtFillSepWords b/tests/bench/fmtFillSepWords new file mode 100644 index 000000000000..5273b6e3f0a2 --- /dev/null +++ b/tests/bench/fmtFillSepWords @@ -0,0 +1,10000 @@ +A +a +aa +aal +aalii +aam +Aani +aardvark +aardwolf +Aaron +Aaronic +Aaronical +Aaronite +Aaronitic +Aaru +Ab +aba +Ababdeh +Ababua +abac +abaca +abacate +abacay +abacinate +abacination +abaciscus +abacist +aback +abactinal +abactinally +abaction +abactor +abaculus +abacus +Abadite +abaff +abaft +abaisance +abaiser +abaissed +abalienate +abalienation +abalone +Abama +abampere +abandon +abandonable +abandoned +abandonedly +abandonee +abandoner +abandonment +Abanic +Abantes +abaptiston +Abarambo +Abaris +abarthrosis +abarticular +abarticulation +abas +abase +abased +abasedly +abasedness +abasement +abaser +Abasgi +abash +abashed +abashedly +abashedness +abashless +abashlessly +abashment +abasia +abasic +abask +Abassin +abastardize +abatable +abate +abatement +abater +abatis +abatised +abaton +abator +abattoir +Abatua +abature +abave +abaxial +abaxile +abaze +abb +Abba +abbacomes +abbacy +Abbadide +abbas +abbasi +abbassi +Abbasside +abbatial +abbatical +abbess +abbey +abbeystede +Abbie +abbot +abbotcy +abbotnullius +abbotship +abbreviate +abbreviately +abbreviation +abbreviator +abbreviatory +abbreviature +Abby +abcoulomb +abdal +abdat +Abderian +Abderite +abdest +abdicable +abdicant +abdicate +abdication +abdicative +abdicator +Abdiel +abditive +abditory +abdomen +abdominal +Abdominales +abdominalian +abdominally +abdominoanterior +abdominocardiac +abdominocentesis +abdominocystic +abdominogenital +abdominohysterectomy +abdominohysterotomy +abdominoposterior +abdominoscope +abdominoscopy +abdominothoracic +abdominous +abdominovaginal +abdominovesical +abduce +abducens +abducent +abduct +abduction +abductor +Abe +abeam +abear +abearance +abecedarian +abecedarium +abecedary +abed +abeigh +Abel +abele +Abelia +Abelian +Abelicea +Abelite +abelite +Abelmoschus +abelmosk +Abelonian +abeltree +Abencerrages +abenteric +abepithymia +Aberdeen +aberdevine +Aberdonian +Aberia +aberrance +aberrancy +aberrant +aberrate +aberration +aberrational +aberrator +aberrometer +aberroscope +aberuncator +abet +abetment +abettal +abettor +abevacuation +abey +abeyance +abeyancy +abeyant +abfarad +abhenry +abhiseka +abhominable +abhor +abhorrence +abhorrency +abhorrent +abhorrently +abhorrer +abhorrible +abhorring +Abhorson +abidal +abidance +abide +abider +abidi +abiding +abidingly +abidingness +Abie +Abies +abietate +abietene +abietic +abietin +Abietineae +abietineous +abietinic +Abiezer +Abigail +abigail +abigailship +abigeat +abigeus +abilao +ability +abilla +abilo +abintestate +abiogenesis +abiogenesist +abiogenetic +abiogenetical +abiogenetically +abiogenist +abiogenous +abiogeny +abiological +abiologically +abiology +abiosis +abiotic +abiotrophic +abiotrophy +Abipon +abir +abirritant +abirritate +abirritation +abirritative +abiston +Abitibi +abiuret +abject +abjectedness +abjection +abjective +abjectly +abjectness +abjoint +abjudge +abjudicate +abjudication +abjunction +abjunctive +abjuration +abjuratory +abjure +abjurement +abjurer +abkar +abkari +Abkhas +Abkhasian +ablach +ablactate +ablactation +ablare +ablastemic +ablastous +ablate +ablation +ablatitious +ablatival +ablative +ablator +ablaut +ablaze +able +ableeze +ablegate +ableness +ablepharia +ablepharon +ablepharous +Ablepharus +ablepsia +ableptical +ableptically +abler +ablest +ablewhackets +ablins +abloom +ablow +ablude +abluent +ablush +ablution +ablutionary +abluvion +ably +abmho +Abnaki +abnegate +abnegation +abnegative +abnegator +Abner +abnerval +abnet +abneural +abnormal +abnormalism +abnormalist +abnormality +abnormalize +abnormally +abnormalness +abnormity +abnormous +abnumerable +Abo +aboard +Abobra +abode +abodement +abody +abohm +aboil +abolish +abolisher +abolishment +abolition +abolitionary +abolitionism +abolitionist +abolitionize +abolla +aboma +abomasum +abomasus +abominable +abominableness +abominably +abominate +abomination +abominator +abomine +Abongo +aboon +aborad +aboral +aborally +abord +aboriginal +aboriginality +aboriginally +aboriginary +aborigine +abort +aborted +aborticide +abortient +abortifacient +abortin +abortion +abortional +abortionist +abortive +abortively +abortiveness +abortus +abouchement +abound +abounder +abounding +aboundingly +about +abouts +above +aboveboard +abovedeck +aboveground +aboveproof +abovestairs +abox +abracadabra +abrachia +abradant +abrade +abrader +Abraham +Abrahamic +Abrahamidae +Abrahamite +Abrahamitic +abraid +Abram +Abramis +abranchial +abranchialism +abranchian +Abranchiata +abranchiate +abranchious +abrasax +abrase +abrash +abrasiometer +abrasion +abrasive +abrastol +abraum +abraxas +abreact +abreaction +abreast +abrenounce +abret +abrico +abridge +abridgeable +abridged +abridgedly +abridger +abridgment +abrim +abrin +abristle +abroach +abroad +Abrocoma +abrocome +abrogable +abrogate +abrogation +abrogative +abrogator +Abroma +Abronia +abrook +abrotanum +abrotine +abrupt +abruptedly +abruption +abruptly +abruptness +Abrus +Absalom +absampere +Absaroka +absarokite +abscess +abscessed +abscession +abscessroot +abscind +abscise +abscision +absciss +abscissa +abscissae +abscisse +abscission +absconce +abscond +absconded +abscondedly +abscondence +absconder +absconsa +abscoulomb +absence +absent +absentation +absentee +absenteeism +absenteeship +absenter +absently +absentment +absentmindedly +absentness +absfarad +abshenry +Absi +absinthe +absinthial +absinthian +absinthiate +absinthic +absinthin +absinthine +absinthism +absinthismic +absinthium +absinthol +absit +absmho +absohm +absolute +absolutely +absoluteness +absolution +absolutism +absolutist +absolutistic +absolutistically +absolutive +absolutization +absolutize +absolutory +absolvable +absolvatory +absolve +absolvent +absolver +absolvitor +absolvitory +absonant +absonous +absorb +absorbability +absorbable +absorbed +absorbedly +absorbedness +absorbefacient +absorbency +absorbent +absorber +absorbing +absorbingly +absorbition +absorpt +absorptance +absorptiometer +absorptiometric +absorption +absorptive +absorptively +absorptiveness +absorptivity +absquatulate +abstain +abstainer +abstainment +abstemious +abstemiously +abstemiousness +abstention +abstentionist +abstentious +absterge +abstergent +abstersion +abstersive +abstersiveness +abstinence +abstinency +abstinent +abstinential +abstinently +abstract +abstracted +abstractedly +abstractedness +abstracter +abstraction +abstractional +abstractionism +abstractionist +abstractitious +abstractive +abstractively +abstractiveness +abstractly +abstractness +abstractor +abstrahent +abstricted +abstriction +abstruse +abstrusely +abstruseness +abstrusion +abstrusity +absume +absumption +absurd +absurdity +absurdly +absurdness +absvolt +Absyrtus +abterminal +abthain +abthainrie +abthainry +abthanage +Abu +abu +abucco +abulia +abulic +abulomania +abuna +abundance +abundancy +abundant +Abundantia +abundantly +abura +aburabozu +aburban +aburst +aburton +abusable +abuse +abusedly +abusee +abuseful +abusefully +abusefulness +abuser +abusion +abusious +abusive +abusively +abusiveness +abut +Abuta +Abutilon +abutment +abuttal +abutter +abutting +abuzz +abvolt +abwab +aby +abysm +abysmal +abysmally +abyss +abyssal +Abyssinian +abyssobenthonic +abyssolith +abyssopelagic +acacatechin +acacatechol +acacetin +Acacia +Acacian +acaciin +acacin +academe +academial +academian +Academic +academic +academical +academically +academicals +academician +academicism +academism +academist +academite +academization +academize +Academus +academy +Acadia +acadialite +Acadian +Acadie +Acaena +acajou +acaleph +Acalepha +Acalephae +acalephan +acalephoid +acalycal +acalycine +acalycinous +acalyculate +Acalypha +Acalypterae +Acalyptrata +Acalyptratae +acalyptrate +Acamar +acampsia +acana +acanaceous +acanonical +acanth +acantha +Acanthaceae +acanthaceous +acanthad +Acantharia +Acanthia +acanthial +acanthin +acanthine +acanthion +acanthite +acanthocarpous +Acanthocephala +acanthocephalan +Acanthocephali +acanthocephalous +Acanthocereus +acanthocladous +Acanthodea +acanthodean +Acanthodei +Acanthodes +acanthodian +Acanthodidae +Acanthodii +Acanthodini +acanthoid +Acantholimon +acanthological +acanthology +acantholysis +acanthoma +Acanthomeridae +acanthon +Acanthopanax +Acanthophis +acanthophorous +acanthopod +acanthopodous +acanthopomatous +acanthopore +acanthopteran +Acanthopteri +acanthopterous +acanthopterygian +Acanthopterygii +acanthosis +acanthous +Acanthuridae +Acanthurus +acanthus +acapnia +acapnial +acapsular +acapu +acapulco +acara +Acarapis +acardia +acardiac +acari +acarian +acariasis +acaricidal +acaricide +acarid +Acarida +Acaridea +acaridean +acaridomatium +acariform +Acarina +acarine +acarinosis +acarocecidium +acarodermatitis +acaroid +acarol +acarologist +acarology +acarophilous +acarophobia +acarotoxic +acarpelous +acarpous +Acarus +Acastus +acatalectic +acatalepsia +acatalepsy +acataleptic +acatallactic +acatamathesia +acataphasia +acataposis +acatastasia +acatastatic +acate +acategorical +acatery +acatharsia +acatharsy +acatholic +acaudal +acaudate +acaulescent +acauline +acaulose +acaulous +acca +accede +accedence +acceder +accelerable +accelerando +accelerant +accelerate +accelerated +acceleratedly +acceleration +accelerative +accelerator +acceleratory +accelerograph +accelerometer +accend +accendibility +accendible +accension +accensor +accent +accentless +accentor +accentuable +accentual +accentuality +accentually +accentuate +accentuation +accentuator +accentus +accept +acceptability +acceptable +acceptableness +acceptably +acceptance +acceptancy +acceptant +acceptation +accepted +acceptedly +accepter +acceptilate +acceptilation +acception +acceptive +acceptor +acceptress +accerse +accersition +accersitor +access +accessarily +accessariness +accessary +accessaryship +accessibility +accessible +accessibly +accession +accessional +accessioner +accessive +accessively +accessless +accessorial +accessorily +accessoriness +accessorius +accessory +accidence +accidency +accident +accidental +accidentalism +accidentalist +accidentality +accidentally +accidentalness +accidented +accidential +accidentiality +accidently +accidia +accidie +accinge +accipient +Accipiter +accipitral +accipitrary +Accipitres +accipitrine +accismus +accite +acclaim +acclaimable +acclaimer +acclamation +acclamator +acclamatory +acclimatable +acclimatation +acclimate +acclimatement +acclimation +acclimatizable +acclimatization +acclimatize +acclimatizer +acclimature +acclinal +acclinate +acclivitous +acclivity +acclivous +accloy +accoast +accoil +accolade +accoladed +accolated +accolent +accolle +accombination +accommodable +accommodableness +accommodate +accommodately +accommodateness +accommodating +accommodatingly +accommodation +accommodational +accommodative +accommodativeness +accommodator +accompanier +accompaniment +accompanimental +accompanist +accompany +accompanyist +accompletive +accomplice +accompliceship +accomplicity +accomplish +accomplishable +accomplished +accomplisher +accomplishment +accomplisht +accompt +accord +accordable +accordance +accordancy +accordant +accordantly +accorder +according +accordingly +accordion +accordionist +accorporate +accorporation +accost +accostable +accosted +accouche +accouchement +accoucheur +accoucheuse +account +accountability +accountable +accountableness +accountably +accountancy +accountant +accountantship +accounting +accountment +accouple +accouplement +accouter +accouterment +accoy +accredit +accreditate +accreditation +accredited +accreditment +accrementitial +accrementition +accresce +accrescence +accrescent +accretal +accrete +accretion +accretionary +accretive +accroach +accroides +accrual +accrue +accruement +accruer +accubation +accubitum +accubitus +accultural +acculturate +acculturation +acculturize +accumbency +accumbent +accumber +accumulable +accumulate +accumulation +accumulativ +accumulative +accumulatively +accumulativeness +accumulator +accuracy +accurate +accurately +accurateness +accurse +accursed +accursedly +accursedness +accusable +accusably +accusal +accusant +accusation +accusatival +accusative +accusatively +accusatorial +accusatorially +accusatory +accusatrix +accuse +accused +accuser +accusingly +accusive +accustom +accustomed +accustomedly +accustomedness +ace +aceacenaphthene +aceanthrene +aceanthrenequinone +acecaffine +aceconitic +acedia +acediamine +acediast +acedy +Aceldama +Acemetae +Acemetic +acenaphthene +acenaphthenyl +acenaphthylene +acentric +acentrous +aceologic +aceology +acephal +Acephala +acephalan +Acephali +acephalia +Acephalina +acephaline +acephalism +acephalist +Acephalite +acephalocyst +acephalous +acephalus +Acer +Aceraceae +aceraceous +Acerae +Acerata +acerate +Acerates +acerathere +Aceratherium +aceratosis +acerb +Acerbas +acerbate +acerbic +acerbity +acerdol +acerin +acerose +acerous +acerra +acertannin +acervate +acervately +acervation +acervative +acervose +acervuline +acervulus +acescence +acescency +acescent +aceship +acesodyne +Acestes +acetabular +Acetabularia +acetabuliferous +acetabuliform +acetabulous +acetabulum +acetacetic +acetal +acetaldehydase +acetaldehyde +acetaldehydrase +acetalization +acetalize +acetamide +acetamidin +acetamidine +acetamido +acetaminol +acetanilid +acetanilide +acetanion +acetaniside +acetanisidide +acetannin +acetarious +acetarsone +acetate +acetated +acetation +acetbromamide +acetenyl +acethydrazide +acetic +acetification +acetifier +acetify +acetimeter +acetimetry +acetin +acetize +acetmethylanilide +acetnaphthalide +acetoacetanilide +acetoacetate +acetoacetic +acetoamidophenol +acetoarsenite +Acetobacter +acetobenzoic +acetobromanilide +acetochloral +acetocinnamene +acetoin +acetol +acetolysis +acetolytic +acetometer +acetometrical +acetometrically +acetometry +acetomorphine +acetonaphthone +acetonate +acetonation +acetone +acetonemia +acetonemic +acetonic +acetonitrile +acetonization +acetonize +acetonuria +acetonurometer +acetonyl +acetonylacetone +acetonylidene +acetophenetide +acetophenin +acetophenine +acetophenone +acetopiperone +acetopyrin +acetosalicylic +acetose +acetosity +acetosoluble +acetothienone +acetotoluide +acetotoluidine +acetous +acetoveratrone +acetoxime +acetoxyl +acetoxyphthalide +acetphenetid +acetphenetidin +acetract +acettoluide +acetum +aceturic +acetyl +acetylacetonates +acetylacetone +acetylamine +acetylate +acetylation +acetylator +acetylbenzene +acetylbenzoate +acetylbenzoic +acetylbiuret +acetylcarbazole +acetylcellulose +acetylcholine +acetylcyanide +acetylenation +acetylene +acetylenediurein +acetylenic +acetylenyl +acetylfluoride +acetylglycine +acetylhydrazine +acetylic +acetylide +acetyliodide +acetylizable +acetylization +acetylize +acetylizer +acetylmethylcarbinol +acetylperoxide +acetylphenol +acetylphenylhydrazine +acetylrosaniline +acetylsalicylate +acetylsalol +acetyltannin +acetylthymol +acetyltropeine +acetylurea +ach +Achaean +Achaemenian +Achaemenid +Achaemenidae +Achaemenidian +Achaenodon +Achaeta +achaetous +achage +Achagua +Achakzai +achalasia +Achamoth +Achango +achar +Achariaceae +Achariaceous +achate +Achates +Achatina +Achatinella +Achatinidae +ache +acheilia +acheilous +acheiria +acheirous +acheirus +Achen +achene +achenial +achenium +achenocarp +achenodium +acher +Achernar +Acheronian +Acherontic +Acherontical +achete +Achetidae +Acheulean +acheweed +achievable +achieve +achievement +achiever +achigan +achilary +achill +Achillea +Achillean +Achilleid +achilleine +Achillize +achillobursitis +achillodynia +achime +Achimenes +Achinese +aching +achingly +achira +Achitophel +achlamydate +Achlamydeae +achlamydeous +achlorhydria +achlorophyllous +achloropsia +Achmetha +acholia +acholic +Acholoe +acholous +acholuria +acholuric +Achomawi +achondrite +achondritic +achondroplasia +achondroplastic +achor +achordal +Achordata +achordate +Achorion +Achras +achree +achroacyte +Achroanthes +achrodextrin +achrodextrinase +achroglobin +achroiocythaemia +achroiocythemia +achroite +achroma +achromacyte +achromasia +achromat +achromate +Achromatiaceae +achromatic +achromatically +achromaticity +achromatin +achromatinic +achromatism +Achromatium +achromatizable +achromatization +achromatize +achromatocyte +achromatolysis +achromatope +achromatophile +achromatopia +achromatopsia +achromatopsy +achromatosis +achromatous +achromaturia +achromia +achromic +Achromobacter +Achromobacterieae +achromoderma +achromophilous +achromotrichia +achromous +achronical +achroodextrin +achroodextrinase +achroous +achropsia +achtehalber +achtel +achtelthaler +Achuas +achy +achylia +achylous +achymia +achymous +Achyranthes +Achyrodes +acichloride +acicula +acicular +acicularly +aciculate +aciculated +aciculum +acid +Acidanthera +Acidaspis +acidemia +acider +acidic +acidiferous +acidifiable +acidifiant +acidific +acidification +acidifier +acidify +acidimeter +acidimetric +acidimetrical +acidimetrically +acidimetry +acidite +acidity +acidize +acidly +acidness +acidoid +acidology +acidometer +acidometry +acidophile +acidophilic +acidophilous +acidoproteolytic +acidosis +acidosteophyte +acidotic +acidproof +acidulate +acidulent +acidulous +aciduric +acidyl +acier +acierage +Acieral +acierate +acieration +aciform +aciliate +aciliated +Acilius +acinaceous +acinaces +acinacifolious +acinaciform +acinar +acinarious +acinary +Acineta +Acinetae +acinetan +Acinetaria +acinetarian +acinetic +acinetiform +Acinetina +acinetinan +acinic +aciniform +acinose +acinotubular +acinous +acinus +Acipenser +Acipenseres +acipenserid +Acipenseridae +acipenserine +acipenseroid +Acipenseroidei +Acis +aciurgy +acker +ackey +ackman +acknow +acknowledge +acknowledgeable +acknowledged +acknowledgedly +acknowledgement +acknowledger +acknowledgment +aclastic +acle +acleidian +acleistous +Aclemon +aclidian +aclinal +aclinic +acloud +aclys +Acmaea +Acmaeidae +acmatic +acme +acmesthesia +acmic +Acmispon +acmite +acne +acneform +acneiform +acnemia +Acnida +acnodal +acnode +Acocanthera +acocantherin +acock +acockbill +acocotl +Acoela +Acoelomata +acoelomate +acoelomatous +Acoelomi +acoelomous +acoelous +Acoemetae +Acoemeti +Acoemetic +acoin +acoine +Acolapissa +acold +Acolhua +Acolhuan +acologic +acology +acolous +acoluthic +acolyte +acolythate +Acoma +acoma +acomia +acomous +aconative +acondylose +acondylous +acone +aconic +aconin +aconine +aconital +aconite +aconitia +aconitic +aconitin +aconitine +Aconitum +Acontias +acontium +Acontius +aconuresis +acopic +acopon +acopyrin +acopyrine +acor +acorea +acoria +acorn +acorned +Acorus +acosmic +acosmism +acosmist +acosmistic +acotyledon +acotyledonous +acouasm +acouchi +acouchy +acoumeter +acoumetry +acouometer +acouophonia +acoupa +acousmata +acousmatic +acoustic +acoustical +acoustically +acoustician +acousticolateral +Acousticon +acoustics +acquaint +acquaintance +acquaintanceship +acquaintancy +acquaintant +acquainted +acquaintedness +acquest +acquiesce +acquiescement +acquiescence +acquiescency +acquiescent +acquiescently +acquiescer +acquiescingly +acquirability +acquirable +acquire +acquired +acquirement +acquirenda +acquirer +acquisible +acquisite +acquisited +acquisition +acquisitive +acquisitively +acquisitiveness +acquisitor +acquisitum +acquist +acquit +acquitment +acquittal +acquittance +acquitter +Acrab +acracy +acraein +Acraeinae +acraldehyde +Acrania +acranial +acraniate +acrasia +Acrasiaceae +Acrasiales +Acrasida +Acrasieae +Acraspeda +acraspedote +acratia +acraturesis +acrawl +acraze +acre +acreable +acreage +acreak +acream +acred +Acredula +acreman +acrestaff +acrid +acridan +acridian +acridic +Acrididae +Acridiidae +acridine +acridinic +acridinium +acridity +Acridium +acridly +acridness +acridone +acridonium +acridophagus +acridyl +acriflavin +acriflavine +acrimonious +acrimoniously +acrimoniousness +acrimony +acrindoline +acrinyl +acrisia +Acrisius +Acrita +acritan +acrite +acritical +acritol +Acroa +acroaesthesia +acroama +acroamatic +acroamatics +acroanesthesia +acroarthritis +acroasphyxia +acroataxia +acroatic +acrobacy +acrobat +Acrobates +acrobatholithic +acrobatic +acrobatical +acrobatically +acrobatics +acrobatism +acroblast +acrobryous +acrobystitis +Acrocarpi +acrocarpous +acrocephalia +acrocephalic +acrocephalous +acrocephaly +Acrocera +Acroceratidae +Acroceraunian +Acroceridae +Acrochordidae +Acrochordinae +acrochordon +Acroclinium +Acrocomia +acroconidium +acrocontracture +acrocoracoid +acrocyanosis +acrocyst +acrodactylum +acrodermatitis +acrodont +acrodontism +acrodrome +acrodromous +Acrodus +acrodynia +acroesthesia +acrogamous +acrogamy +acrogen +acrogenic +acrogenous +acrogenously +acrography +Acrogynae +acrogynae +acrogynous +acrolein +acrolith +acrolithan +acrolithic +acrologic +acrologically +acrologism +acrologue +acrology +acromania +acromastitis +acromegalia +acromegalic +acromegaly +acromelalgia +acrometer +acromial +acromicria +acromioclavicular +acromiocoracoid +acromiodeltoid +acromiohumeral +acromiohyoid +acromion +acromioscapular +acromiosternal +acromiothoracic +acromonogrammatic +acromphalus +Acromyodi +acromyodian +acromyodic +acromyodous +acromyotonia +acromyotonus +acron +acronarcotic +acroneurosis +acronical +acronically +acronyc +acronych +Acronycta +acronyctous +acronym +acronymic +acronymize +acronymous +acronyx +acrook +acroparalysis +acroparesthesia +acropathology +acropathy +acropetal +acropetally +acrophobia +acrophonetic +acrophonic +acrophony +acropodium +acropoleis +acropolis +acropolitan +Acropora +acrorhagus +acrorrheuma +acrosarc +acrosarcum +acroscleriasis +acroscleroderma +acroscopic +acrose +acrosome +acrosphacelus +acrospire +acrospore +acrosporous +across +acrostic +acrostical +acrostically +acrostichal +Acrosticheae +acrostichic +acrostichoid +Acrostichum +acrosticism +acrostolion +acrostolium +acrotarsial +acrotarsium +acroteleutic +acroterial +acroteric +acroterium +Acrothoracica +acrotic +acrotism +acrotomous +Acrotreta +Acrotretidae +acrotrophic +acrotrophoneurosis +Acrux +Acrydium +acryl +acrylaldehyde +acrylate +acrylic +acrylonitrile +acrylyl +act +acta +actability +actable +Actaea +Actaeaceae +Actaeon +Actaeonidae +Actiad +Actian +actification +actifier +actify +actin +actinal +actinally +actinautographic +actinautography +actine +actinenchyma +acting +Actinia +actinian +Actiniaria +actiniarian +actinic +actinically +Actinidia +Actinidiaceae +actiniferous +actiniform +actinine +actiniochrome +actiniohematin +Actiniomorpha +actinism +Actinistia +actinium +actinobacillosis +Actinobacillus +actinoblast +actinobranch +actinobranchia +actinocarp +actinocarpic +actinocarpous +actinochemistry +actinocrinid +Actinocrinidae +actinocrinite +Actinocrinus +actinocutitis +actinodermatitis +actinodielectric +actinodrome +actinodromous +actinoelectric +actinoelectrically +actinoelectricity +actinogonidiate +actinogram +actinograph +actinography +actinoid +Actinoida +Actinoidea +actinolite +actinolitic +actinologous +actinologue +actinology +actinomere +actinomeric +actinometer +actinometric +actinometrical +actinometry +actinomorphic +actinomorphous +actinomorphy +Actinomyces +Actinomycetaceae +Actinomycetales +actinomycete +actinomycetous +actinomycin +actinomycoma +actinomycosis +actinomycotic +Actinomyxidia +Actinomyxidiida +actinon +Actinonema +actinoneuritis +actinophone +actinophonic +actinophore +actinophorous +actinophryan +Actinophrys +Actinopoda +actinopraxis +actinopteran +Actinopteri +actinopterous +actinopterygian +Actinopterygii +actinopterygious +actinoscopy +actinosoma +actinosome +Actinosphaerium +actinost +actinostereoscopy +actinostomal +actinostome +actinotherapeutic +actinotherapeutics +actinotherapy +actinotoxemia +actinotrichium +actinotrocha +actinouranium +Actinozoa +actinozoal +actinozoan +actinozoon +actinula +action +actionable +actionably +actional +actionary +actioner +actionize +actionless +Actipylea +Actium +activable +activate +activation +activator +active +actively +activeness +activin +activism +activist +activital +activity +activize +actless +actomyosin +acton +actor +actorship +actress +Acts +actu +actual +actualism +actualist +actualistic +actuality +actualization +actualize +actually +actualness +actuarial +actuarially +actuarian +actuary +actuaryship +actuation +actuator +acture +acturience +actutate +acuaesthesia +Acuan +acuate +acuation +Acubens +acuclosure +acuductor +acuesthesia +acuity +aculea +Aculeata +aculeate +aculeated +aculeiform +aculeolate +aculeolus +aculeus +acumen +acuminate +acumination +acuminose +acuminous +acuminulate +acupress +acupressure +acupunctuate +acupunctuation +acupuncturation +acupuncturator +acupuncture +acurative +acushla +acutangular +acutate +acute +acutely +acutenaculum +acuteness +acutiator +acutifoliate +Acutilinguae +acutilingual +acutilobate +acutiplantar +acutish +acutograve +acutonodose +acutorsion +acyanoblepsia +acyanopsia +acyclic +acyesis +acyetic +acyl +acylamido +acylamidobenzene +acylamino +acylate +acylation +acylogen +acyloin +acyloxy +acyloxymethane +acyrological +acyrology +acystia +ad +Ada +adactyl +adactylia +adactylism +adactylous +Adad +adad +adage +adagial +adagietto +adagio +Adai +Adaize +Adam +adamant +adamantean +adamantine +adamantinoma +adamantoblast +adamantoblastoma +adamantoid +adamantoma +adamas +Adamastor +adambulacral +adamellite +Adamhood +Adamic +Adamical +Adamically +adamine +Adamite +adamite +Adamitic +Adamitical +Adamitism +Adamsia +adamsite +adance +adangle +Adansonia +Adapa +adapid +Adapis +adapt +adaptability +adaptable +adaptation +adaptational +adaptationally +adaptative +adaptedness +adapter +adaption +adaptional +adaptionism +adaptitude +adaptive +adaptively +adaptiveness +adaptometer +adaptor +adaptorial +Adar +adarme +adat +adati +adatom +adaunt +adaw +adawe +adawlut +adawn +adaxial +aday +adays +adazzle +adcraft +add +Adda +adda +addability +addable +addax +addebted +added +addedly +addend +addenda +addendum +adder +adderbolt +adderfish +adderspit +adderwort +addibility +addible +addicent +addict +addicted +addictedness +addiction +Addie +addiment +Addisonian +Addisoniana +additament +additamentary +addition +additional +additionally +additionary +additionist +addititious +additive +additively +additivity +additory +addle +addlebrain +addlebrained +addlehead +addleheaded +addleheadedly +addleheadedness +addlement +addleness +addlepate +addlepated +addlepatedness +addleplot +addlings +addlins +addorsed +address +addressee +addresser +addressful +Addressograph +addressor +addrest +Addu +adduce +adducent +adducer +adducible +adduct +adduction +adductive +adductor +Addy +Ade +ade +adead +adeem +adeep +Adela +Adelaide +Adelarthra +Adelarthrosomata +adelarthrosomatous +Adelbert +Adelea +Adeleidae +Adelges +Adelia +Adelina +Adeline +adeling +adelite +Adeliza +adelocerous +Adelochorda +adelocodonic +adelomorphic +adelomorphous +adelopod +Adelops +Adelphi +Adelphian +adelphogamy +Adelphoi +adelpholite +adelphophagy +ademonist +adempted +ademption +adenalgia +adenalgy +Adenanthera +adenase +adenasthenia +adendric +adendritic +adenectomy +adenectopia +adenectopic +adenemphractic +adenemphraxis +adenia +adeniform +adenine +adenitis +adenization +adenoacanthoma +adenoblast +adenocancroid +adenocarcinoma +adenocarcinomatous +adenocele +adenocellulitis +adenochondroma +adenochondrosarcoma +adenochrome +adenocyst +adenocystoma +adenocystomatous +adenodermia +adenodiastasis +adenodynia +adenofibroma +adenofibrosis +adenogenesis +adenogenous +adenographer +adenographic +adenographical +adenography +adenohypersthenia +adenoid +adenoidal +adenoidism +adenoliomyofibroma +adenolipoma +adenolipomatosis +adenologaditis +adenological +adenology +adenolymphocele +adenolymphoma +adenoma +adenomalacia +adenomatome +adenomatous +adenomeningeal +adenometritis +adenomycosis +adenomyofibroma +adenomyoma +adenomyxoma +adenomyxosarcoma +adenoncus +adenoneural +adenoneure +adenopathy +adenopharyngeal +adenopharyngitis +adenophlegmon +Adenophora +adenophore +adenophorous +adenophthalmia +adenophyllous +adenophyma +adenopodous +adenosarcoma +adenosclerosis +adenose +adenosine +adenosis +adenostemonous +Adenostoma +adenotome +adenotomic +adenotomy +adenotyphoid +adenotyphus +adenyl +adenylic +Adeodatus +Adeona +Adephaga +adephagan +adephagia +adephagous +adept +adeptness +adeptship +adequacy +adequate +adequately +adequateness +adequation +adequative +adermia +adermin +Adessenarian +adet +adevism +adfected +adfix +adfluxion +adglutinate +Adhafera +adhaka +adhamant +Adhara +adharma +adhere +adherence +adherency +adherent +adherently +adherer +adherescence +adherescent +adhesion +adhesional +adhesive +adhesively +adhesivemeter +adhesiveness +adhibit +adhibition +adiabatic +adiabatically +adiabolist +adiactinic +adiadochokinesis +adiagnostic +adiantiform +Adiantum +adiaphon +adiaphonon +adiaphoral +adiaphoresis +adiaphoretic +adiaphorism +adiaphorist +adiaphoristic +adiaphorite +adiaphoron +adiaphorous +adiate +adiathermal +adiathermancy +adiathermanous +adiathermic +adiathetic +adiation +Adib +Adicea +adicity +Adiel +adieu +adieux +Adigei +Adighe +Adigranth +adigranth +Adin +Adinida +adinidan +adinole +adion +adipate +adipescent +adipic +adipinic +adipocele +adipocellulose +adipocere +adipoceriform +adipocerous +adipocyte +adipofibroma +adipogenic +adipogenous +adipoid +adipolysis +adipolytic +adipoma +adipomatous +adipometer +adipopexia +adipopexis +adipose +adiposeness +adiposis +adiposity +adiposogenital +adiposuria +adipous +adipsia +adipsic +adipsous +adipsy +adipyl +Adirondack +adit +adital +aditus +adjacency +adjacent +adjacently +adjag +adject +adjection +adjectional +adjectival +adjectivally +adjective +adjectively +adjectivism +adjectivitis +adjiger +adjoin +adjoined +adjoinedly +adjoining +adjoint +adjourn +adjournal +adjournment +adjudge +adjudgeable +adjudger +adjudgment +adjudicate +adjudication +adjudicative +adjudicator +adjudicature +adjunct +adjunction +adjunctive +adjunctively +adjunctly +adjuration +adjuratory +adjure +adjurer +adjust +adjustable +adjustably +adjustage +adjustation +adjuster +adjustive +adjustment +adjutage +adjutancy +adjutant +adjutantship +adjutorious +adjutory +adjutrice +adjuvant +Adlai +adlay +adless +adlet +Adlumia +adlumidine +adlumine +adman +admarginate +admaxillary +admeasure +admeasurement +admeasurer +admedial +admedian +admensuration +admi +admin +adminicle +adminicula +adminicular +adminiculary +adminiculate +adminiculation +adminiculum +administer +administered +administerial +administrable +administrant +administrate +administration +administrational +administrative +administratively +administrator +administratorship +administratress +administratrices +administratrix +admirability +admirable +admirableness +admirably +admiral +admiralship +admiralty +admiration +admirative +admirator +admire +admired +admiredly +admirer +admiring +admiringly +admissibility +admissible +admissibleness +admissibly +admission +admissive +admissory +admit +admittable +admittance +admitted +admittedly +admittee +admitter +admittible +admix +admixtion +admixture +admonish +admonisher +admonishingly +admonishment +admonition +admonitioner +admonitionist +admonitive +admonitively +admonitor +admonitorial +admonitorily +admonitory +admonitrix +admortization +adnascence +adnascent +adnate +adnation +adnephrine +adnerval +adneural +adnex +adnexal +adnexed +adnexitis +adnexopexy +adnominal +adnominally +adnomination +adnoun +ado +adobe +adolesce +adolescence +adolescency +adolescent +adolescently +Adolph +Adolphus +Adonai +Adonean +Adonia +Adoniad +Adonian +Adonic +adonidin +adonin +Adoniram +Adonis +adonite +adonitol +adonize +adoperate +adoperation +adopt +adoptability +adoptable +adoptant +adoptative +adopted +adoptedly +adoptee +adopter +adoptian +adoptianism +adoptianist +adoption +adoptional +adoptionism +adoptionist +adoptious +adoptive +adoptively +adorability +adorable +adorableness +adorably +adoral +adorally +adorant +Adorantes +adoration +adoratory +adore +adorer +Adoretus +adoringly +adorn +adorner +adorningly +adornment +adosculation +adossed +adoulie +adown +Adoxa +Adoxaceae +adoxaceous +adoxography +adoxy +adoze +adpao +adpress +adpromission +adradial +adradially +adradius +Adramelech +Adrammelech +adread +adream +adreamed +adreamt +adrectal +adrenal +adrenalectomize +adrenalectomy +Adrenalin +adrenaline +adrenalize +adrenalone +adrenergic +adrenin +adrenine +adrenochrome +adrenocortical +adrenocorticotropic +adrenolysis +adrenolytic +adrenotropic +Adrian +Adriana +Adriatic +Adrienne +adrift +adrip +adroit +adroitly +adroitness +adroop +adrop +adrostral +adrowse +adrue +adry +adsbud +adscendent +adscititious +adscititiously +adscript +adscripted +adscription +adscriptitious +adscriptitius +adscriptive +adsessor +adsheart +adsignification +adsignify +adsmith +adsmithing +adsorb +adsorbable +adsorbate +adsorbent +adsorption +adsorptive +adstipulate +adstipulation +adstipulator +adterminal +adtevac +adular +adularescence +adularia +adulate +adulation +adulator +adulatory +adulatress +Adullam +Adullamite +adult +adulter +adulterant +adulterate +adulterately +adulterateness +adulteration +adulterator +adulterer +adulteress +adulterine +adulterize +adulterous +adulterously +adultery +adulthood +adulticidal +adulticide +adultness +adultoid +adumbral +adumbrant +adumbrate +adumbration +adumbrative +adumbratively +adunc +aduncate +aduncated +aduncity +aduncous +adusk +adust +adustion +adustiosis +Advaita +advance +advanceable +advanced +advancedness +advancement +advancer +advancing +advancingly +advancive +advantage +advantageous +advantageously +advantageousness +advection +advectitious +advective +advehent +advene +advenience +advenient +Advent +advential +Adventism +Adventist +adventitia +adventitious +adventitiously +adventitiousness +adventive +adventual +adventure +adventureful +adventurement +adventurer +adventureship +adventuresome +adventuresomely +adventuresomeness +adventuress +adventurish +adventurous +adventurously +adventurousness +adverb +adverbial +adverbiality +adverbialize +adverbially +adverbiation +adversant +adversaria +adversarious +adversary +adversative +adversatively +adverse +adversely +adverseness +adversifoliate +adversifolious +adversity +advert +advertence +advertency +advertent +advertently +advertisable +advertise +advertisee +advertisement +advertiser +advertising +advice +adviceful +advisability +advisable +advisableness +advisably +advisal +advisatory +advise +advised +advisedly +advisedness +advisee +advisement +adviser +advisership +advisive +advisiveness +advisor +advisorily +advisory +advocacy +advocate +advocateship +advocatess +advocation +advocator +advocatory +advocatress +advocatrice +advocatrix +advolution +advowee +advowson +ady +adynamia +adynamic +adynamy +adyta +adyton +adytum +adz +adze +adzer +adzooks +ae +Aeacides +Aeacus +Aeaean +Aechmophorus +aecial +Aecidiaceae +aecidial +aecidioform +Aecidiomycetes +aecidiospore +aecidiostage +aecidium +aeciospore +aeciostage +aecioteliospore +aeciotelium +aecium +aedeagus +Aedes +aedicula +aedile +aedileship +aedilian +aedilic +aedilitian +aedility +aedoeagus +aefald +aefaldness +aefaldy +aefauld +aegagropila +aegagropile +aegagrus +Aegean +aegerian +aegeriid +Aegeriidae +Aegialitis +aegicrania +Aegina +Aeginetan +Aeginetic +Aegipan +aegirine +aegirinolite +aegirite +aegis +Aegisthus +Aegithalos +Aegithognathae +aegithognathism +aegithognathous +Aegle +Aegopodium +aegrotant +aegyptilla +aegyrite +aeluroid +Aeluroidea +aelurophobe +aelurophobia +aeluropodous +aenach +aenean +aeneolithic +aeneous +aenigmatite +aeolharmonica +Aeolia +Aeolian +Aeolic +Aeolicism +aeolid +Aeolidae +Aeolididae +aeolina +aeoline +aeolipile +Aeolis +Aeolism +Aeolist +aeolistic +aeolodicon +aeolodion +aeolomelodicon +aeolopantalon +aeolotropic +aeolotropism +aeolotropy +aeolsklavier +aeon +aeonial +aeonian +aeonist +Aepyceros +Aepyornis +Aepyornithidae +Aepyornithiformes +Aequi +Aequian +Aequiculi +Aequipalpia +aequoreal +aer +aerage +aerarian +aerarium +aerate +aeration +aerator +aerenchyma +aerenterectasia +aerial +aerialist +aeriality +aerially +aerialness +aeric +aerical +Aerides +aerie +aeried +aerifaction +aeriferous +aerification +aeriform +aerify +aero +Aerobacter +aerobate +aerobatic +aerobatics +aerobe +aerobian +aerobic +aerobically +aerobiologic +aerobiological +aerobiologically +aerobiologist +aerobiology +aerobion +aerobiont +aerobioscope +aerobiosis +aerobiotic +aerobiotically +aerobious +aerobium +aeroboat +Aerobranchia +aerobranchiate +aerobus +aerocamera +aerocartograph +Aerocharidae +aerocolpos +aerocraft +aerocurve +aerocyst +aerodermectasia +aerodone +aerodonetic +aerodonetics +aerodrome +aerodromics +aerodynamic +aerodynamical +aerodynamicist +aerodynamics +aerodyne +aeroembolism +aeroenterectasia +aerofoil +aerogel +aerogen +aerogenes +aerogenesis +aerogenic +aerogenically +aerogenous +aerogeologist +aerogeology +aerognosy +aerogram +aerograph +aerographer +aerographic +aerographical +aerographics +aerography +aerogun +aerohydrodynamic +aerohydropathy +aerohydroplane +aerohydrotherapy +aerohydrous +aeroides +aerolite +aerolith +aerolithology +aerolitic +aerolitics +aerologic +aerological +aerologist +aerology +aeromaechanic +aeromancer +aeromancy +aeromantic +aeromarine +aeromechanical +aeromechanics +aerometeorograph +aerometer +aerometric +aerometry +aeromotor +aeronat +aeronaut +aeronautic +aeronautical +aeronautically +aeronautics +aeronautism +aeronef +aeroneurosis +aeropathy +Aerope +aeroperitoneum +aeroperitonia +aerophagia +aerophagist +aerophagy +aerophane +aerophilatelic +aerophilatelist +aerophilately +aerophile +aerophilic +aerophilous +aerophobia +aerophobic +aerophone +aerophor +aerophore +aerophotography +aerophysical +aerophysics +aerophyte +aeroplane +aeroplaner +aeroplanist +aeropleustic +aeroporotomy +aeroscepsis +aeroscepsy +aeroscope +aeroscopic +aeroscopically +aeroscopy +aerose +aerosiderite +aerosiderolite +Aerosol +aerosol +aerosphere +aerosporin +aerostat +aerostatic +aerostatical +aerostatics +aerostation +aerosteam +aerotactic +aerotaxis +aerotechnical +aerotherapeutics +aerotherapy +aerotonometer +aerotonometric +aerotonometry +aerotropic +aerotropism +aeroyacht +aeruginous +aerugo +aery +aes +Aeschylean +Aeschynanthus +Aeschynomene +aeschynomenous +Aesculaceae +aesculaceous +Aesculapian +Aesculapius +Aesculus +Aesopian +Aesopic +aesthete +aesthetic +aesthetical +aesthetically +aesthetician +aestheticism +aestheticist +aestheticize +aesthetics +aesthiology +aesthophysiology +Aestii +aethalioid +aethalium +aetheogam +aetheogamic +aetheogamous +aethered +Aethionema +aethogen +aethrioscope +Aethusa +Aetian +aetiogenic +aetiotropic +aetiotropically +Aetobatidae +Aetobatus +Aetolian +Aetomorphae +aetosaur +aetosaurian +Aetosaurus +aevia +aface +afaint +Afar +afar +afara +afear +afeard +afeared +afebrile +Afenil +afernan +afetal +affa +affability +affable +affableness +affably +affabrous +affair +affaite +affect +affectable +affectate +affectation +affectationist +affected +affectedly +affectedness +affecter +affectibility +affectible +affecting +affectingly +affection +affectional +affectionally +affectionate +affectionately +affectionateness +affectioned +affectious +affective +affectively +affectivity +affeer +affeerer +affeerment +affeir +affenpinscher +affenspalte +afferent +affettuoso +affiance +affiancer +affiant +affidation +affidavit +affidavy +affiliable +affiliate +affiliation +affinal +affination +affine +affined +affinely +affinitative +affinitatively +affinite +affinition +affinitive +affinity +affirm +affirmable +affirmably +affirmance +affirmant +affirmation +affirmative +affirmatively +affirmatory +affirmer +affirmingly +affix +affixal +affixation +affixer +affixion +affixture +afflation +afflatus +afflict +afflicted +afflictedness +afflicter +afflicting +afflictingly +affliction +afflictionless +afflictive +afflictively +affluence +affluent +affluently +affluentness +afflux +affluxion +afforce +afforcement +afford +affordable +afforest +afforestable +afforestation +afforestment +afformative +affranchise +affranchisement +affray +affrayer +affreight +affreighter +affreightment +affricate +affricated +affrication +affricative +affright +affrighted +affrightedly +affrighter +affrightful +affrightfully +affrightingly +affrightment +affront +affronte +affronted +affrontedly +affrontedness +affronter +affronting +affrontingly +affrontingness +affrontive +affrontiveness +affrontment +affuse +affusion +affy +Afghan +afghani +afield +Afifi +afikomen +afire +aflagellar +aflame +aflare +aflat +aflaunt +aflicker +aflight +afloat +aflow +aflower +afluking +aflush +aflutter +afoam +afoot +afore +aforehand +aforenamed +aforesaid +aforethought +aforetime +aforetimes +afortiori +afoul +afraid +afraidness +Aframerican +Afrasia +Afrasian +afreet +afresh +afret +Afric +African +Africana +Africanism +Africanist +Africanization +Africanize +Africanoid +Africanthropus +Afridi +Afrikaans +Afrikander +Afrikanderdom +Afrikanderism +Afrikaner +Afrogaea +Afrogaean +afront +afrown +Afshah +Afshar +aft +aftaba +after +afteract +afterage +afterattack +afterband +afterbeat +afterbirth +afterblow +afterbody +afterbrain +afterbreach +afterbreast +afterburner +afterburning +aftercare +aftercareer +aftercast +aftercataract +aftercause +afterchance +afterchrome +afterchurch +afterclap +afterclause +aftercome +aftercomer +aftercoming +aftercooler +aftercost +aftercourse +aftercrop +aftercure +afterdamp +afterdate +afterdays +afterdeck +afterdinner +afterdrain +afterdrops +aftereffect +afterend +aftereye +afterfall +afterfame +afterfeed +afterfermentation +afterform +afterfriend +afterfruits +afterfuture +aftergame +aftergas +afterglide +afterglow +aftergo +aftergood +aftergrass +aftergrave +aftergrief +aftergrind +aftergrowth +afterguard +afterguns +afterhand +afterharm +afterhatch +afterhelp +afterhend +afterhold +afterhope +afterhours +afterimage +afterimpression +afterings +afterking +afterknowledge +afterlife +afterlifetime +afterlight +afterloss +afterlove +aftermark +aftermarriage +aftermass +aftermast +aftermath +aftermatter +aftermeal +aftermilk +aftermost +afternight +afternoon +afternoons +afternose +afternote +afteroar +afterpain +afterpart +afterpast +afterpeak +afterpiece +afterplanting +afterplay +afterpressure +afterproof +afterrake +afterreckoning +afterrider +afterripening +afterroll +afterschool +aftersend +aftersensation +aftershaft +aftershafted +aftershine +aftership +aftershock +aftersong +aftersound +afterspeech +afterspring +afterstain +afterstate +afterstorm +afterstrain +afterstretch +afterstudy +afterswarm +afterswarming +afterswell +aftertan +aftertask +aftertaste +afterthinker +afterthought +afterthoughted +afterthrift +aftertime +aftertimes +aftertouch +aftertreatment +aftertrial +afterturn +aftervision +afterwale +afterwar +afterward +afterwards +afterwash +afterwhile +afterwisdom +afterwise +afterwit +afterwitted +afterwork +afterworking +afterworld +afterwrath +afterwrist +aftmost +Aftonian +aftosa +aftward +aftwards +afunction +afunctional +afwillite +Afzelia +aga +agabanee +agacante +agacella +Agaces +Agade +Agag +again +against +againstand +agal +agalactia +agalactic +agalactous +agalawood +agalaxia +agalaxy +Agalena +Agalenidae +Agalinis +agalite +agalloch +agallochum +agallop +agalma +agalmatolite +agalwood +Agama +agama +Agamae +Agamemnon +agamete +agami +agamian +agamic +agamically +agamid +Agamidae +agamobium +agamogenesis +agamogenetic +agamogenetically +agamogony +agamoid +agamont +agamospore +agamous +agamy +aganglionic +Aganice +Aganippe +Agao +Agaonidae +Agapanthus +agape +Agapemone +Agapemonian +Agapemonist +Agapemonite +agapetae +agapeti +agapetid +Agapetidae +Agapornis +agar +agaric +agaricaceae +agaricaceous +Agaricales +agaricic +agariciform +agaricin +agaricine +agaricoid +Agaricus +Agaristidae +agarita +Agarum +agarwal +agasp +Agastache +Agastreae +agastric +agastroneuria +agate +agateware +Agatha +Agathaea +Agathaumas +agathin +Agathis +agathism +agathist +agathodaemon +agathodaemonic +agathokakological +agathology +Agathosma +agatiferous +agatiform +agatine +agatize +agatoid +agaty +Agau +Agave +agavose +Agawam +Agaz +agaze +agazed +Agdistis +age +aged +agedly +agedness +agee +Agelacrinites +Agelacrinitidae +Agelaius +agelast +Agelaus +ageless +agelessness +agelong +agen +Agena +agency +agenda +agendum +agenesia +agenesic +agenesis +agennetic +agent +agentess +agential +agentival +agentive +agentry +agentship +ageometrical +ager +Ageratum +ageusia +ageusic +ageustia +agger +aggerate +aggeration +aggerose +Aggie +agglomerant +agglomerate +agglomerated +agglomeratic +agglomeration +agglomerative +agglomerator +agglutinability +agglutinable +agglutinant +agglutinate +agglutination +agglutinationist +agglutinative +agglutinator +agglutinin +agglutinize +agglutinogen +agglutinogenic +agglutinoid +agglutinoscope +agglutogenic +aggradation +aggradational +aggrade +aggrandizable +aggrandize +aggrandizement +aggrandizer +aggrate +aggravate +aggravating +aggravatingly +aggravation +aggravative +aggravator +aggregable +aggregant +Aggregata +Aggregatae +aggregate +aggregately +aggregateness +aggregation +aggregative +aggregator +aggregatory +aggress +aggressin +aggression +aggressionist +aggressive +aggressively +aggressiveness +aggressor +aggrievance +aggrieve +aggrieved +aggrievedly +aggrievedness +aggrievement +aggroup +aggroupment +aggry +aggur +agha +Aghan +aghanee +aghast +aghastness +Aghlabite +Aghorapanthi +Aghori +Agialid +Agib +Agiel +agilawood +agile +agilely +agileness +agility +agillawood +aging +agio +agiotage +agist +agistator +agistment +agistor +agita +agitable +agitant +agitate +agitatedly +agitation +agitational +agitationist +agitative +agitator +agitatorial +agitatrix +agitprop +Agkistrodon +agla +Aglaia +aglance +Aglaonema +Aglaos +aglaozonia +aglare +Aglaspis +Aglauros +agleaf +agleam +aglet +aglethead +agley +aglimmer +aglint +Aglipayan +Aglipayano +aglitter +aglobulia +Aglossa +aglossal +aglossate +aglossia +aglow +aglucon +aglutition +aglycosuric +Aglypha +aglyphodont +Aglyphodonta +Aglyphodontia +aglyphous +agmatine +agmatology +agminate +agminated +agnail +agname +agnamed +agnate +Agnatha +agnathia +agnathic +Agnathostomata +agnathostomatous +agnathous +agnatic +agnatically +agnation +agnel +Agnes +agnification +agnize +Agnoetae +Agnoete +Agnoetism +agnoiology +Agnoite +agnomen +agnomical +agnominal +agnomination +agnosia +agnosis +agnostic +agnostically +agnosticism +Agnostus +agnosy +Agnotozoic +agnus +ago +agog +agoge +agogic +agogics +agoho +agoing +agomensin +agomphiasis +agomphious +agomphosis +agon +agonal +agone +agoniada +agoniadin +agoniatite +Agoniatites +agonic +agonied +agonist +Agonista +agonistarch +agonistic +agonistically +agonistics +agonium +agonize +agonizedly +agonizer +agonizingly +Agonostomus +agonothete +agonothetic +agony +agora +agoranome +agoraphobia +agouara +agouta +agouti +agpaite +agpaitic +Agra +agraffee +agrah +agral +agrammatical +agrammatism +Agrania +agranulocyte +agranulocytosis +agranuloplastic +Agrapha +agraphia +agraphic +agrarian +agrarianism +agrarianize +agrarianly +Agrauleum +agre +agree +agreeability +agreeable +agreeableness +agreeably +agreed +agreeing +agreeingly +agreement +agreer +agregation +agrege +agrestal +agrestial +agrestian +agrestic +agria +agricere +agricole +agricolist +agricolite +agricolous +agricultor +agricultural +agriculturalist +agriculturally +agriculture +agriculturer +agriculturist +Agrilus +Agrimonia +agrimony +agrimotor +agrin +Agriochoeridae +Agriochoerus +agriological +agriologist +agriology +Agrionia +agrionid +Agrionidae +Agriotes +Agriotypidae +Agriotypus +agrise +agrito +agroan +agrobiologic +agrobiological +agrobiologically +agrobiologist +agrobiology +agrogeological +agrogeologically +agrogeology +agrologic +agrological +agrologically +agrology +agrom +Agromyza +agromyzid +Agromyzidae +agronome +agronomial +agronomic +agronomical +agronomics +agronomist +agronomy +agroof +agrope +Agropyron +Agrostemma +agrosteral +Agrostis +agrostographer +agrostographic +agrostographical +agrostography +agrostologic +agrostological +agrostologist +agrostology +agrotechny +Agrotis +aground +agrufe +agruif +agrypnia +agrypnotic +agsam +agua +aguacate +Aguacateca +aguavina +Agudist +ague +aguelike +agueproof +agueweed +aguey +aguilarite +aguilawood +aguinaldo +aguirage +aguish +aguishly +aguishness +agunah +agush +agust +agy +Agyieus +agynarious +agynary +agynous +agyrate +agyria +Ah +ah +aha +ahaaina +ahankara +Ahantchuyuk +ahartalav +ahaunch +ahead +aheap +ahem +Ahepatokla +Ahet +ahey +ahimsa +ahind +ahint +Ahir +ahluwalia +ahmadi +Ahmadiya +Ahmed +Ahmet +Ahnfeltia +aho +Ahom +ahong +ahorse +ahorseback +Ahousaht +ahoy +Ahrendahronon +Ahriman +Ahrimanian +ahsan +Aht +Ahtena +ahu +ahuatle +ahuehuete +ahull +ahum +ahungered +ahungry +ahunt +ahura +ahush +ahwal +ahypnia +ai +Aias +Aiawong +aichmophobia +aid +aidable +aidance +aidant +aide +Aidenn +aider +Aides +aidful +aidless +aiel +aigialosaur +Aigialosauridae +Aigialosaurus +aiglet +aigremore +aigrette +aiguille +aiguillesque +aiguillette +aiguilletted +aikinite +ail +ailantery +ailanthic +Ailanthus +ailantine +ailanto +aile +Aileen +aileron +ailette +Ailie +ailing +aillt +ailment +ailsyte +Ailuridae +ailuro +ailuroid +Ailuroidea +Ailuropoda +Ailuropus +Ailurus +ailweed +aim +Aimak +aimara +Aimee +aimer +aimful +aimfully +aiming +aimless +aimlessly +aimlessness +Aimore +aimworthiness +ainaleh +ainhum +ainoi +ainsell +aint +Ainu +aion +aionial +air +Aira +airable +airampo +airan +airbound +airbrained +airbrush +aircraft +aircraftman +aircraftsman +aircraftswoman +aircraftwoman +aircrew +aircrewman +airdock +airdrome +airdrop +aire +Airedale +airedale +airer +airfield +airfoil +airframe +airfreight +airfreighter +airgraphics +airhead +airiferous +airified +airily +airiness +airing +airish +airless +airlift +airlike +airliner +airmail +airman +airmanship +airmark +airmarker +airmonger +airohydrogen +airometer +airpark +airphobia +airplane +airplanist +airport +airproof +airscape +airscrew +airship +airsick +airsickness +airstrip +airt +airtight +airtightly +airtightness +airward +airwards +airway +airwayman +airwoman +airworthiness +airworthy +airy +aischrolatreia +aiseweed +aisle +aisled +aisleless +aisling +Aissaoua +Aissor +aisteoir +Aistopoda +Aistopodes +ait +aitch +aitchbone +aitchless +aitchpiece +aitesis +aithochroi +aition +aitiotropic +Aitkenite +Aitutakian +aiwan +Aix +aizle +Aizoaceae +aizoaceous +Aizoon +Ajaja +ajaja +ajangle +ajar +ajari +Ajatasatru +ajava +ajhar +ajivika +ajog +ajoint +ajowan +Ajuga +ajutment +ak +Aka +aka +Akal +akala +Akali +akalimba +akamatsu +Akamnik +Akan +Akanekunik +Akania +Akaniaceae +akaroa +akasa +Akawai +akazga +akazgine +akcheh +ake +akeake +akebi +Akebia +akee +akeki +akeley +akenobeite +akepiro +akerite +akey +Akha +Akhissar +Akhlame +Akhmimic +akhoond +akhrot +akhyana +akia +Akim +akimbo +akin +akindle +akinesia +akinesic +akinesis +akinete +akinetic +Akiskemikinik +Akiyenik +Akka +Akkad +Akkadian +Akkadist +akmudar +akmuddar +aknee +ako +akoasm +akoasma +akoluthia +akonge +Akontae +Akoulalion +akov +akpek +Akra +akra +Akrabattine +akroasis +akrochordite +akroterion +Aktistetae +Aktistete +Aktivismus +Aktivist +aku +akuammine +akule +akund +Akwapim +Al +al +ala +Alabama +Alabaman +Alabamian +alabamide +alabamine +alabandite +alabarch +alabaster +alabastos +alabastrian +alabastrine +alabastrites +alabastron +alabastrum +alacha +alack +alackaday +alacreatine +alacreatinine +alacrify +alacritous +alacrity +Alactaga +alada +Aladdin +Aladdinize +Aladfar +Aladinist +alaihi +Alain +alaite +Alaki +Alala +alala +alalite +alalonga +alalunga +alalus +Alamanni +Alamannian +Alamannic +alameda +alamo +alamodality +alamonti +alamosite +alamoth +Alan +alan +aland +Alangiaceae +alangin +alangine +Alangium +alani +alanine +alannah +Alans +alantic +alantin +alantol +alantolactone +alantolic +alanyl +alar +Alarbus +alares +Alaria +Alaric +alarm +alarmable +alarmed +alarmedly +alarming +alarmingly +alarmism +alarmist +Alarodian +alarum +alary +alas +Alascan +Alaska +alaskaite +Alaskan +alaskite +Alastair +Alaster +alastrim +alate +alated +alatern +alaternus +alation +Alauda +Alaudidae +alaudine +Alaunian +Alawi +Alb +alb +alba +albacore +albahaca +Albainn +Alban +alban +Albanenses +Albanensian +Albania +Albanian +albanite +Albany +albarco +albardine +albarello +albarium +albaspidin +albata +Albatros +albatross +albe +albedo +albedograph +albee +albeit +Alberene +Albert +Alberta +albertin +Albertina +Albertine +Albertinian +Albertist +albertite +Alberto +albertustaler +albertype +albescence +albescent +albespine +albetad +Albi +Albian +albicans +albicant +albication +albiculi +albification +albificative +albiflorous +albify +Albigenses +Albigensian +Albigensianism +Albin +albinal +albiness +albinic +albinism +albinistic +albino +albinoism +albinotic +albinuria +Albion +Albireo +albite +albitic +albitite +albitization +albitophyre +Albizzia +albocarbon +albocinereous +Albococcus +albocracy +Alboin +albolite +albolith +albopannin +albopruinose +alboranite +Albrecht +Albright +albronze +Albruna +Albuca +Albuginaceae +albuginea +albugineous +albuginitis +albugo +album +albumean +albumen +albumenization +albumenize +albumenizer +albumimeter +albumin +albuminate +albuminaturia +albuminiferous +albuminiform +albuminimeter +albuminimetry +albuminiparous +albuminization +albuminize +albuminocholia +albuminofibrin +albuminogenous +albuminoid +albuminoidal +albuminolysis +albuminometer +albuminometry +albuminone +albuminorrhea +albuminoscope +albuminose +albuminosis +albuminous +albuminousness +albuminuria +albuminuric +albumoid +albumoscope +albumose +albumosuria +alburn +alburnous +alburnum +albus +albutannin +Albyn +Alca +Alcaaba +Alcae +Alcaic +alcaide +alcalde +alcaldeship +alcaldia +Alcaligenes +alcalizate +Alcalzar +alcamine +alcanna +Alcantara +Alcantarines +alcarraza +alcatras +alcazar +Alcedines +Alcedinidae +Alcedininae +Alcedo +alcelaphine +Alcelaphus +Alces +alchemic +alchemical +alchemically +Alchemilla +alchemist +alchemistic +alchemistical +alchemistry +alchemize +alchemy +alchera +alcheringa +alchimy +alchitran +alchochoden +Alchornea +alchymy +Alcibiadean +Alcicornium +Alcidae +alcidine +alcine +Alcippe +alclad +alco +alcoate +alcogel +alcogene +alcohate +alcohol +alcoholate +alcoholature +alcoholdom +alcoholemia +alcoholic +alcoholically +alcoholicity +alcoholimeter +alcoholism +alcoholist +alcoholizable +alcoholization +alcoholize +alcoholmeter +alcoholmetric +alcoholomania +alcoholometer +alcoholometric +alcoholometrical +alcoholometry +alcoholophilia +alcoholuria +alcoholysis +alcoholytic +Alcor +Alcoran +Alcoranic +Alcoranist +alcornoco +alcornoque +alcosol +Alcotate +alcove +alcovinometer +Alcuinian +alcyon +Alcyonacea +alcyonacean +Alcyonaria +alcyonarian +Alcyone +Alcyones +Alcyoniaceae +alcyonic +alcyoniform +Alcyonium +alcyonoid +aldamine +aldane +aldazin +aldazine +aldeament +Aldebaran +aldebaranium +aldehol +aldehydase +aldehyde +aldehydic +aldehydine +aldehydrol +alder +Alderamin +alderman +aldermanate +aldermancy +aldermaness +aldermanic +aldermanical +aldermanity +aldermanlike +aldermanly +aldermanry +aldermanship +aldern +Alderney +alderwoman +Aldhafara +Aldhafera +aldim +aldime +aldimine +Aldine +aldine +aldoheptose +aldohexose +aldoketene +aldol +aldolization +aldolize +aldononose +aldopentose +aldose +aldoside +aldoxime +Aldrovanda +Aldus +ale +Alea +aleak +aleatory +alebench +aleberry +Alebion +alec +alecithal +alecize +Aleck +aleconner +alecost +Alectoria +alectoria +Alectorides +alectoridine +alectorioid +Alectoris +alectoromachy +alectoromancy +Alectoromorphae +alectoromorphous +Alectoropodes +alectoropodous +Alectrion +Alectrionidae +alectryomachy +alectryomancy +Alectryon +alecup +alee +alef +alefnull +aleft +alefzero +alegar +alehoof +alehouse +Alejandro +alem +alemana +Alemanni +Alemannian +Alemannic +Alemannish +alembic +alembicate +alembroth +Alemite +alemite +alemmal +alemonger +alen +Alencon +Aleochara +aleph +alephs +alephzero +alepidote +alepole +alepot +Aleppine +Aleppo +alerce +alerse +alert +alertly +alertness +alesan +alestake +aletap +aletaster +Alethea +alethiology +alethopteis +alethopteroid +alethoscope +aletocyte +Aletris +alette +aleukemic +Aleurites +aleuritic +Aleurobius +Aleurodes +Aleurodidae +aleuromancy +aleurometer +aleuronat +aleurone +aleuronic +aleuroscope +Aleut +Aleutian +Aleutic +aleutite +alevin +alewife +Alex +Alexander +alexanders +Alexandra +Alexandreid +Alexandrian +Alexandrianism +Alexandrina +Alexandrine +alexandrite +Alexas +Alexia +alexia +Alexian +alexic +alexin +alexinic +alexipharmacon +alexipharmacum +alexipharmic +alexipharmical +alexipyretic +Alexis +alexiteric +alexiterical +Alexius +aleyard +Aleyrodes +aleyrodid +Aleyrodidae +Alf +alf +alfa +alfaje +alfalfa +alfaqui +alfaquin +alfenide +alfet +alfilaria +alfileria +alfilerilla +alfilerillo +alfiona +Alfirk +alfonsin +alfonso +alforja +Alfred +Alfreda +alfresco +alfridaric +alfridary +Alfur +Alfurese +Alfuro +alga +algae +algaecide +algaeological +algaeologist +algaeology +algaesthesia +algaesthesis +algal +algalia +Algaroth +algarroba +algarrobilla +algarrobin +Algarsife +Algarsyf +algate +Algebar +algebra +algebraic +algebraical +algebraically +algebraist +algebraization +algebraize +Algedi +algedo +algedonic +algedonics +algefacient +Algenib +Algerian +Algerine +algerine +Algernon +algesia +algesic +algesis +algesthesis +algetic +Algic +algic +algid +algidity +algidness +Algieba +algific +algin +alginate +algine +alginic +alginuresis +algiomuscular +algist +algivorous +algocyan +algodoncillo +algodonite +algoesthesiometer +algogenic +algoid +Algol +algolagnia +algolagnic +algolagnist +algolagny +algological +algologist +algology +Algoman +algometer +algometric +algometrical +algometrically +algometry +Algomian +Algomic +Algonkian +Algonquian +Algonquin +algophilia +algophilist +algophobia +algor +Algorab +Algores +algorism +algorismic +algorist +algoristic +algorithm +algorithmic +algosis +algous +algovite +algraphic +algraphy +alguazil +algum +Algy +Alhagi +Alhambra +Alhambraic +Alhambresque +Alhena +alhenna +alias +Alibamu +alibangbang +alibi +alibility +alible +Alicant +Alice +alichel +Alichino +Alicia +Alick +alicoche +alictisal +alicyclic +Alida +alidade +Alids +alien +alienability +alienable +alienage +alienate +alienation +alienator +aliency +alienee +aliener +alienicola +alienigenate +alienism +alienist +alienize +alienor +alienship +aliethmoid +aliethmoidal +alif +aliferous +aliform +aligerous +alight +align +aligner +alignment +aligreek +aliipoe +alike +alikeness +alikewise +Alikuluf +Alikulufan +alilonghi +alima +aliment +alimental +alimentally +alimentariness +alimentary +alimentation +alimentative +alimentatively +alimentativeness +alimenter +alimentic +alimentive +alimentiveness +alimentotherapy +alimentum +alimonied +alimony +alin +alinasal +Aline +alineation +alintatao +aliofar +Alioth +alipata +aliped +aliphatic +alipterion +aliptes +aliptic +aliquant +aliquot +aliseptal +alish +alisier +Alisma +Alismaceae +alismaceous +alismad +alismal +Alismales +Alismataceae +alismoid +aliso +Alison +alison +alisonite +alisp +alisphenoid +alisphenoidal +alist +Alister +alit +alite +alitrunk +aliturgic +aliturgical +aliunde +alive +aliveness +alivincular +Alix +aliyah +alizarate +alizari +alizarin +aljoba +alk +alkahest +alkahestic +alkahestica +alkahestical +Alkaid +alkalamide +alkalemia +alkalescence +alkalescency +alkalescent +alkali +alkalic +alkaliferous +alkalifiable +alkalify +alkaligen +alkaligenous +alkalimeter +alkalimetric +alkalimetrical +alkalimetrically +alkalimetry +alkaline +alkalinity +alkalinization +alkalinize +alkalinuria +alkalizable +alkalizate +alkalization +alkalize +alkalizer +alkaloid +alkaloidal +alkalometry +alkalosis +alkalous +Alkalurops +alkamin +alkamine +alkane +alkanet +Alkanna +alkannin +Alkaphrah +alkapton +alkaptonuria +alkaptonuric +alkargen +alkarsin +alkekengi +alkene +alkenna +alkenyl +alkermes +Alkes +alkide +alkine +alkool +Alkoran +Alkoranic +alkoxide +alkoxy +alkoxyl +alky +alkyd +alkyl +alkylamine +alkylate +alkylation +alkylene +alkylic +alkylidene +alkylize +alkylogen +alkyloxy +alkyne +all +allabuta +allactite +allaeanthus +allagite +allagophyllous +allagostemonous +Allah +allalinite +Allamanda +allamotti +Allan +allan +allanite +allanitic +allantiasis +allantochorion +allantoic +allantoid +allantoidal +Allantoidea +allantoidean +allantoidian +allantoin +allantoinase +allantoinuria +allantois +allantoxaidin +allanturic +Allasch +allassotonic +allative +allatrate +allay +allayer +allayment +allbone +Alle +allecret +allectory +allegate +allegation +allegator +allege +allegeable +allegedly +allegement +alleger +Alleghenian +Allegheny +allegiance +allegiancy +allegiant +allegoric +allegorical +allegorically +allegoricalness +allegorism +allegorist +allegorister +allegoristic +allegorization +allegorize +allegorizer +allegory +allegretto +allegro +allele +allelic +allelism +allelocatalytic +allelomorph +allelomorphic +allelomorphism +allelotropic +allelotropism +allelotropy +alleluia +alleluiatic +allemand +allemande +allemontite +Allen +allenarly +allene +Allentiac +Allentiacan +aller +allergen +allergenic +allergia +allergic +allergin +allergist +allergy +allerion +allesthesia +alleviate +alleviatingly +alleviation +alleviative +alleviator +alleviatory +alley +alleyed +alleyite +alleyway +allgood +Allhallow +Allhallowtide +allheal +alliable +alliably +Alliaceae +alliaceous +alliance +alliancer +Alliaria +allicampane +allice +allicholly +alliciency +allicient +Allie +allied +Allies +allies +alligate +alligator +alligatored +allineate +allineation +Allionia +Allioniaceae +allision +alliteral +alliterate +alliteration +alliterational +alliterationist +alliterative +alliteratively +alliterativeness +alliterator +Allium +allivalite +allmouth +allness +Allobroges +allocable +allocaffeine +allocatable +allocate +allocatee +allocation +allocator +allochetia +allochetite +allochezia +allochiral +allochirally +allochiria +allochlorophyll +allochroic +allochroite +allochromatic +allochroous +allochthonous +allocinnamic +alloclase +alloclasite +allocochick +allocrotonic +allocryptic +allocute +allocution +allocutive +allocyanine +allodelphite +allodesmism +alloeosis +alloeostropha +alloeotic +alloerotic +alloerotism +allogamous +allogamy +allogene +allogeneity +allogeneous +allogenic +allogenically +allograph +alloiogenesis +alloisomer +alloisomeric +alloisomerism +allokinesis +allokinetic +allokurtic +allomerism +allomerous +allometric +allometry +allomorph +allomorphic +allomorphism +allomorphite +allomucic +allonomous +allonym +allonymous +allopalladium +allopath +allopathetic +allopathetically +allopathic +allopathically +allopathist +allopathy +allopatric +allopatrically +allopatry +allopelagic +allophanamide +allophanates +allophane +allophanic +allophone +allophyle +allophylian +allophylic +Allophylus +allophytoid +alloplasm +alloplasmatic +alloplasmic +alloplast +alloplastic +alloplasty +alloploidy +allopolyploid +allopsychic +alloquial +alloquialism +alloquy +allorhythmia +allorrhyhmia +allorrhythmic +allosaur +Allosaurus +allose +allosematic +allosome +allosyndesis +allosyndetic +allot +allotee +allotelluric +allotheism +Allotheria +allothigene +allothigenetic +allothigenetically +allothigenic +allothigenous +allothimorph +allothimorphic +allothogenic +allothogenous +allotment +allotriodontia +Allotriognathi +allotriomorphic +allotriophagia +allotriophagy +allotriuria +allotrope +allotrophic +allotropic +allotropical +allotropically +allotropicity +allotropism +allotropize +allotropous +allotropy +allotrylic +allottable +allottee +allotter +allotype +allotypical +allover +allow +allowable +allowableness +allowably +allowance +allowedly +allower +alloxan +alloxanate +alloxanic +alloxantin +alloxuraemia +alloxuremia +alloxuric +alloxyproteic +alloy +alloyage +allozooid +allseed +allspice +allthing +allthorn +alltud +allude +allure +allurement +allurer +alluring +alluringly +alluringness +allusion +allusive +allusively +allusiveness +alluvia +alluvial +alluviate +alluviation +alluvion +alluvious +alluvium +allwhere +allwhither +allwork +Allworthy +Ally +ally +allyl +allylamine +allylate +allylation +allylene +allylic +allylthiourea +Alma +alma +Almach +almaciga +almacigo +almadia +almadie +almagest +almagra +Almain +Alman +almanac +almandine +almandite +alme +almeidina +almemar +Almerian +almeriite +Almida +almightily +almightiness +almighty +almique +Almira +almirah +almochoden +Almohad +Almohade +Almohades +almoign +Almon +almon +almond +almondy +almoner +almonership +almonry +Almoravid +Almoravide +Almoravides +almost +almous +alms +almsdeed +almsfolk +almsful +almsgiver +almsgiving +almshouse +almsman +almswoman +almucantar +almuce +almud +almude +almug +Almuredin +almuten +aln +alnage +alnager +alnagership +Alnaschar +Alnascharism +alnein +alnico +Alnilam +alniresinol +Alnitak +Alnitham +alniviridol +alnoite +alnuin +Alnus +alo +Aloadae +Alocasia +alochia +alod +alodial +alodialism +alodialist +alodiality +alodially +alodian +alodiary +alodification +alodium +alody +aloe +aloed +aloelike +aloemodin +aloeroot +aloesol +aloeswood +aloetic +aloetical +aloewood +aloft +alogia +Alogian +alogical +alogically +alogism +alogy +aloid +aloin +Alois +aloisiite +aloma +alomancy +alone +aloneness +along +alongshore +alongshoreman +alongside +alongst +Alonso +Alonsoa +Alonzo +aloof +aloofly +aloofness +aloose +alop +alopecia +Alopecias +alopecist +alopecoid +Alopecurus +alopeke +Alopias +Alopiidae +Alosa +alose +Alouatta +alouatte +aloud +alow +alowe +Aloxite +Aloysia +Aloysius +alp +alpaca +alpasotes +Alpax +alpeen +Alpen +alpenglow +alpenhorn +alpenstock +alpenstocker +alpestral +alpestrian +alpestrine +alpha +alphabet +alphabetarian +alphabetic +alphabetical +alphabetically +alphabetics +alphabetiform +alphabetism +alphabetist +alphabetization +alphabetize +alphabetizer +alphanumeric +Alphard +alphatoluic +Alphean +Alphecca +alphenic +Alpheratz +alphitomancy +alphitomorphous +alphol +Alphonist +Alphonse +Alphonsine +Alphonsism +Alphonso +alphorn +alphos +alphosis +alphyl +Alpian +Alpid +alpieu +alpigene +Alpine +alpine +alpinely +alpinery +alpinesque +Alpinia +Alpiniaceae +Alpinism +Alpinist +alpist +Alpujarra +alqueire +alquier +alquifou +alraun +alreadiness +already +alright +alrighty +alroot +alruna +Alsatia +Alsatian +alsbachite +Alshain +Alsinaceae +alsinaceous +Alsine +also +alsoon +Alsophila +Alstonia +alstonidine +alstonine +alstonite +Alstroemeria +alsweill +alt +Altaian +Altaic +Altaid +Altair +altaite +Altamira +altar +altarage +altared +altarist +altarlet +altarpiece +altarwise +altazimuth +alter +alterability +alterable +alterableness +alterably +alterant +alterate +alteration +alterative +altercate +altercation +altercative +alteregoism +alteregoistic +alterer +alterity +altern +alternacy +alternance +alternant +Alternanthera +Alternaria +alternariose +alternate +alternately +alternateness +alternating +alternatingly +alternation +alternationist +alternative +alternatively +alternativeness +alternativity +alternator +alterne +alternifoliate +alternipetalous +alternipinnate +alternisepalous +alternize +alterocentric +Althaea +althaein +Althea +althea +althein +altheine +althionic +altho +althorn +although +Altica +Alticamelus +altigraph +altilik +altiloquence +altiloquent +altimeter +altimetrical +altimetrically +altimetry +altin +altincar +Altingiaceae +altingiaceous +altininck +altiplano +altiscope +altisonant +altisonous +altissimo +altitude +altitudinal +altitudinarian +alto +altogether +altogetherness +altometer +altoun +altrices +altricial +altropathy +altrose +altruism +altruist +altruistic +altruistically +altschin +altun +Aluco +Aluconidae +Aluconinae +aludel +Aludra +alula +alular +alulet +Alulim +alum +alumbloom +Alumel +alumic +alumiferous +alumina +aluminaphone +aluminate +alumine +aluminic +aluminide +aluminiferous +aluminiform +aluminish +aluminite +aluminium +aluminize +aluminoferric +aluminographic +aluminography +aluminose +aluminosilicate +aluminosis +aluminosity +aluminothermic +aluminothermics +aluminothermy +aluminotype +aluminous +aluminum +aluminyl +alumish +alumite +alumium +alumna +alumnae +alumnal +alumni +alumniate +Alumnol +alumnus +alumohydrocalcite +alumroot +Alundum +aluniferous +alunite +alunogen +alupag +Alur +alure +alurgite +alushtite +aluta +alutaceous +Alvah +Alvan +alvar +alvearium +alveary +alveloz +alveola +alveolar +alveolariform +alveolary +alveolate +alveolated +alveolation +alveole +alveolectomy +alveoli +alveoliform +alveolite +Alveolites +alveolitis +alveoloclasia +alveolocondylean +alveolodental +alveololabial +alveololingual +alveolonasal +alveolosubnasal +alveolotomy +alveolus +alveus +alviducous +Alvin +Alvina +alvine +Alvissmal +alvite +alvus +alway +always +aly +Alya +alycompaine +alymphia +alymphopotent +alypin +alysson +Alyssum +alytarch +Alytes +am +ama +amaas +Amabel +amability +amacratic +amacrinal +amacrine +amadavat +amadelphous +Amadi +Amadis +amadou +Amaethon +Amafingo +amaga +amah +Amahuaca +amain +amaister +amakebe +Amakosa +amala +amalaita +amalaka +Amalfian +Amalfitan +amalgam +amalgamable +amalgamate +amalgamation +amalgamationist +amalgamative +amalgamatize +amalgamator +amalgamist +amalgamization +amalgamize +Amalings +Amalrician +amaltas +amamau +Amampondo +Amanda +amandin +Amandus +amang +amani +amania +Amanist +Amanita +amanitin +amanitine +Amanitopsis +amanori +amanous +amantillo +amanuenses +amanuensis +amapa +Amapondo +amar +Amara +Amarantaceae +amarantaceous +amaranth +Amaranthaceae +amaranthaceous +amaranthine +amaranthoid +Amaranthus +amarantite +Amarantus +amarelle +amarevole +amargoso +amarillo +amarin +amarine +amaritude +amarity +amaroid +amaroidal +Amarth +amarthritis +amaryllid +Amaryllidaceae +amaryllidaceous +amaryllideous +Amaryllis +amasesis +amass +amassable +amasser +amassment +Amasta +amasthenic +amastia +amasty +Amatembu +amaterialistic +amateur +amateurish +amateurishly +amateurishness +amateurism +amateurship +Amati +amative +amatively +amativeness +amatol +amatorial +amatorially +amatorian +amatorious +amatory +amatrice +amatungula +amaurosis +amaurotic +amaze +amazed +amazedly +amazedness +amazeful +amazement +amazia +Amazilia +amazing +amazingly +Amazon +Amazona +Amazonian +Amazonism +amazonite +Amazulu +amba +ambage +ambagiosity +ambagious +ambagiously +ambagiousness +ambagitory +ambalam +amban +ambar +ambaree +ambarella +ambary +ambash +ambassade +Ambassadeur +ambassador +ambassadorial +ambassadorially +ambassadorship +ambassadress +ambassage +ambassy +ambatch +ambatoarinite +ambay +ambeer +amber +amberfish +ambergris +amberiferous +amberite +amberoid +amberous +ambery +ambicolorate +ambicoloration +ambidexter +ambidexterity +ambidextral +ambidextrous +ambidextrously +ambidextrousness +ambience +ambiency +ambiens +ambient +ambier +ambigenous +ambiguity +ambiguous +ambiguously +ambiguousness +ambilateral +ambilateralaterally +ambilaterality +ambilevous +ambilian +ambilogy +ambiopia +ambiparous +ambisinister +ambisinistrous +ambisporangiate +ambisyllabic +ambit +ambital +ambitendency +ambition +ambitionist +ambitionless +ambitionlessly +ambitious +ambitiously +ambitiousness +ambitty +ambitus +ambivalence +ambivalency +ambivalent +ambivert +amble +ambler +ambling +amblingly +amblotic +amblyacousia +amblyaphia +Amblycephalidae +Amblycephalus +amblychromatic +Amblydactyla +amblygeusia +amblygon +amblygonal +amblygonite +amblyocarpous +Amblyomma +amblyope +amblyopia +amblyopic +Amblyopsidae +Amblyopsis +amblyoscope +amblypod +Amblypoda +amblypodous +Amblyrhynchus +amblystegite +Amblystoma +ambo +amboceptoid +amboceptor +Ambocoelia +Amboina +Amboinese +ambomalleal +ambon +ambonite +Ambonnay +ambos +ambosexous +ambosexual +ambrain +ambrein +ambrette +Ambrica +ambrite +ambroid +ambrology +Ambrose +ambrose +ambrosia +ambrosiac +Ambrosiaceae +ambrosiaceous +ambrosial +ambrosially +Ambrosian +ambrosian +ambrosiate +ambrosin +ambrosine +Ambrosio +ambrosterol +ambrotype +ambry +ambsace +ambulacral +ambulacriform +ambulacrum +ambulance +ambulancer +ambulant +ambulate +ambulatio +ambulation +ambulative +ambulator +Ambulatoria +ambulatorial +ambulatorium +ambulatory +ambuling +ambulomancy +amburbial +ambury +ambuscade +ambuscader +ambush +ambusher +ambushment +Ambystoma +Ambystomidae +amchoor +ame +amebiform +Amedeo +ameed +ameen +Ameiuridae +Ameiurus +Ameiva +Amelanchier +amelcorn +Amelia +amelia +amelification +ameliorable +ameliorableness +ameliorant +ameliorate +amelioration +ameliorativ +ameliorative +ameliorator +amellus +ameloblast +ameloblastic +amelu +amelus +Amen +amen +amenability +amenable +amenableness +amenably +amend +amendable +amendableness +amendatory +amende +amender +amendment +amends +amene +amenia +Amenism +Amenite +amenity +amenorrhea +amenorrheal +amenorrheic +amenorrhoea +ament +amentaceous +amental +amentia +Amentiferae +amentiferous +amentiform +amentulum +amentum +amerce +amerceable +amercement +amercer +amerciament +America +American +Americana +Americanese +Americanism +Americanist +Americanistic +Americanitis +Americanization +Americanize +Americanizer +Americanly +Americanoid +Americaward +Americawards +americium +Americomania +Americophobe +Amerimnon +Amerind +Amerindian +Amerindic +amerism +ameristic +amesite +Ametabola +ametabole +ametabolia +ametabolian +ametabolic +ametabolism +ametabolous +ametaboly +ametallous +amethodical +amethodically +amethyst +amethystine +ametoecious +ametria +ametrometer +ametrope +ametropia +ametropic +ametrous +Amex +amgarn +amhar +amherstite +amhran +Ami +ami +Amia +amiability +amiable +amiableness +amiably +amianth +amianthiform +amianthine +Amianthium +amianthoid +amianthoidal +amianthus +amic +amicability +amicable +amicableness +amicably +amical +amice +amiced +amicicide +amicrobic +amicron +amicronucleate +amid +amidase +amidate +amidation +amide +amidic +amidid +amidide +amidin +amidine +Amidism +Amidist +amido +amidoacetal +amidoacetic +amidoacetophenone +amidoaldehyde +amidoazo +amidoazobenzene +amidoazobenzol +amidocaffeine +amidocapric +amidofluorid +amidofluoride +amidogen +amidoguaiacol +amidohexose +amidoketone +amidol +amidomyelin +amidon +amidophenol +amidophosphoric +amidoplast +amidoplastid +amidopyrine +amidosuccinamic +amidosulphonal +amidothiazole +amidoxime +amidoxy +amidoxyl +amidrazone +amidship +amidships +amidst +amidstream +amidulin +Amigo +Amiidae +amil +Amiles +Amiloun +amimia +amimide +amin +aminate +amination +amine +amini +aminic +aminity +aminization +aminize +amino +aminoacetal +aminoacetanilide +aminoacetic +aminoacetone +aminoacetophenetidine +aminoacetophenone +aminoacidemia +aminoaciduria +aminoanthraquinone +aminoazobenzene +aminobarbituric +aminobenzaldehyde +aminobenzamide +aminobenzene +aminobenzoic +aminocaproic +aminodiphenyl +aminoethionic +aminoformic +aminogen +aminoglutaric +aminoguanidine +aminoid +aminoketone +aminolipin +aminolysis +aminolytic +aminomalonic +aminomyelin +aminophenol +aminoplast +aminoplastic +aminopropionic +aminopurine +aminopyrine +aminoquinoline +aminosis +aminosuccinamic +aminosulphonic +aminothiophen +aminovaleric +aminoxylol +Aminta +Amintor +Amioidei +Amir +amir +Amiranha +amiray +amirship +Amish +Amishgo +amiss +amissibility +amissible +amissness +Amita +Amitabha +amitosis +amitotic +amitotically +amity +amixia +Amizilis +amla +amli +amlikar +amlong +Amma +amma +amman +Ammanite +ammelide +ammelin +ammeline +ammer +ammeter +Ammi +Ammiaceae +ammiaceous +ammine +amminochloride +amminolysis +amminolytic +ammiolite +ammo +Ammobium +ammochaeta +ammochryse +ammocoete +ammocoetes +ammocoetid +Ammocoetidae +ammocoetiform +ammocoetoid +Ammodytes +Ammodytidae +ammodytoid +ammonal +ammonate +ammonation +Ammonea +ammonia +ammoniacal +ammoniacum +ammoniate +ammoniation +ammonic +ammonical +ammoniemia +ammonification +ammonifier +ammonify +ammoniojarosite +ammonion +ammonionitrate +Ammonite +ammonite +Ammonites +Ammonitess +ammonitic +ammoniticone +ammonitiferous +Ammonitish +ammonitoid +Ammonitoidea +ammonium +ammoniuria +ammonization +ammono +ammonobasic +ammonocarbonic +ammonocarbonous +ammonoid +Ammonoidea +ammonoidean +ammonolysis +ammonolytic +ammonolyze +Ammophila +ammophilous +ammoresinol +ammotherapy +ammu +ammunition +amnemonic +amnesia +amnesic +amnestic +amnesty +amnia +amniac +amniatic +amnic +Amnigenia +amnioallantoic +amniocentesis +amniochorial +amnioclepsis +amniomancy +amnion +Amnionata +amnionate +amnionic +amniorrhea +Amniota +amniote +amniotic +amniotitis +amniotome +amober +amobyr +amoeba +amoebae +Amoebaea +amoebaean +amoebaeum +amoebalike +amoeban +amoebian +amoebiasis +amoebic +amoebicide +amoebid +Amoebida +Amoebidae +amoebiform +Amoebobacter +Amoebobacterieae +amoebocyte +Amoebogeniae +amoeboid +amoeboidism +amoebous +amoebula +amok +amoke +amole +amolilla +amomal +Amomales +Amomis +amomum +among +amongst +amontillado +amor +amorado +amoraic +amoraim +amoral +amoralism +amoralist +amorality +amoralize +Amores +amoret +amoretto +Amoreuxia +amorism +amorist +amoristic +Amorite +Amoritic +Amoritish +amorosity +amoroso +amorous +amorously +amorousness +Amorpha +amorphia +amorphic +amorphinism +amorphism +Amorphophallus +amorphophyte +amorphotae +amorphous +amorphously +amorphousness +amorphus +amorphy +amort +amortisseur +amortizable +amortization +amortize +amortizement +Amorua +Amos +Amoskeag +amotion +amotus +amount +amour +amourette +amovability +amovable +amove +Amoy +Amoyan +Amoyese +amp +ampalaya +ampalea +ampangabeite +ampasimenite +Ampelidaceae +ampelidaceous +Ampelidae +ampelideous +Ampelis +ampelite +ampelitic +ampelographist +ampelography +ampelopsidin +ampelopsin +Ampelopsis +Ampelosicyos +ampelotherapy +amper +amperage +ampere +amperemeter +Amperian +amperometer +ampersand +ampery +amphanthium +ampheclexis +ampherotokous +ampherotoky +amphetamine +amphiarthrodial +amphiarthrosis +amphiaster +amphibalus +Amphibia +amphibial +amphibian +amphibichnite +amphibiety +amphibiological +amphibiology +amphibion +amphibiotic +Amphibiotica +amphibious +amphibiously +amphibiousness +amphibium +amphiblastic +amphiblastula +amphiblestritis +Amphibola +amphibole +amphibolia +amphibolic +amphiboliferous +amphiboline +amphibolite +amphibolitic +amphibological +amphibologically +amphibologism +amphibology +amphibolous +amphiboly +amphibrach +amphibrachic +amphibryous +Amphicarpa +Amphicarpaea +amphicarpic +amphicarpium +amphicarpogenous +amphicarpous +amphicentric +amphichroic +amphichrom +amphichromatic +amphichrome +amphicoelian +amphicoelous +Amphicondyla +amphicondylous +amphicrania +amphicreatinine +amphicribral +amphictyon +amphictyonian +amphictyonic +amphictyony +Amphicyon +Amphicyonidae +amphicyrtic +amphicyrtous +amphicytula +amphid +amphide +amphidesmous +amphidetic +amphidiarthrosis +amphidiploid +amphidiploidy +amphidisc +Amphidiscophora +amphidiscophoran +amphierotic +amphierotism +Amphigaea +amphigam +Amphigamae +amphigamous +amphigastrium +amphigastrula +amphigean +amphigen +amphigene +amphigenesis +amphigenetic +amphigenous +amphigenously +amphigonic +amphigonium +amphigonous +amphigony +amphigoric +amphigory +amphigouri +amphikaryon +amphilogism +amphilogy +amphimacer +amphimictic +amphimictical +amphimictically +amphimixis +amphimorula +Amphinesian +Amphineura +amphineurous +amphinucleus +Amphion +Amphionic +Amphioxi +Amphioxidae +Amphioxides +Amphioxididae +amphioxus +amphipeptone +amphiphloic +amphiplatyan +Amphipleura +amphiploid +amphiploidy +amphipneust +Amphipneusta +amphipneustic +Amphipnous +amphipod +Amphipoda +amphipodal +amphipodan +amphipodiform +amphipodous +amphiprostylar +amphiprostyle +amphiprotic +amphipyrenin +Amphirhina +amphirhinal +amphirhine +amphisarca +amphisbaena +amphisbaenian +amphisbaenic +Amphisbaenidae +amphisbaenoid +amphisbaenous +amphiscians +amphiscii +Amphisile +Amphisilidae +amphispermous +amphisporangiate +amphispore +Amphistoma +amphistomatic +amphistome +amphistomoid +amphistomous +Amphistomum +amphistylar +amphistylic +amphistyly +amphitene +amphitheater +amphitheatered +amphitheatral +amphitheatric +amphitheatrical +amphitheatrically +amphithecial +amphithecium +amphithect +amphithyron +amphitokal +amphitokous +amphitoky +amphitriaene +amphitrichous +Amphitrite +amphitropal +amphitropous +Amphitruo +Amphitryon +Amphiuma +Amphiumidae +amphivasal +amphivorous +Amphizoidae +amphodarch +amphodelite +amphodiplopia +amphogenous +ampholyte +amphopeptone +amphophil +amphophile +amphophilic +amphophilous +amphora +amphoral +amphore +amphorette +amphoric +amphoricity +amphoriloquy +amphorophony +amphorous +amphoteric +Amphrysian +ample +amplectant +ampleness +amplexation +amplexicaudate +amplexicaul +amplexicauline +amplexifoliate +amplexus +ampliate +ampliation +ampliative +amplicative +amplidyne +amplification +amplificative +amplificator +amplificatory +amplifier +amplify +amplitude +amply +ampollosity +ampongue +ampoule +ampul +ampulla +ampullaceous +ampullar +Ampullaria +Ampullariidae +ampullary +ampullate +ampullated +ampulliform +ampullitis +ampullula +amputate +amputation +amputational +amputative +amputator +amputee +ampyx +amra +amreeta +amrita +Amritsar +amsath +amsel +Amsonia +Amsterdamer +amt +amtman +Amuchco +amuck +Amueixa +amuguis +amula +amulet +amuletic +amulla +amunam +amurca +amurcosity +amurcous +Amurru +amusable +amuse +amused +amusedly +amusee +amusement +amuser +amusette +Amusgo +amusia +amusing +amusingly +amusingness +amusive +amusively +amusiveness +amutter +amuyon +amuyong +amuze +amvis +Amy +amy +Amyclaean +Amyclas +amyelencephalia +amyelencephalic +amyelencephalous +amyelia +amyelic +amyelinic +amyelonic +amyelous +amygdal +amygdala +Amygdalaceae +amygdalaceous +amygdalase +amygdalate +amygdalectomy +amygdalic +amygdaliferous +amygdaliform +amygdalin +amygdaline +amygdalinic +amygdalitis +amygdaloid +amygdaloidal +amygdalolith +amygdaloncus +amygdalopathy +amygdalothripsis +amygdalotome +amygdalotomy +Amygdalus +amygdonitrile +amygdophenin +amygdule +amyl +amylaceous +amylamine +amylan +amylase +amylate +amylemia +amylene +amylenol +amylic +amylidene +amyliferous +amylin +amylo +amylocellulose +amyloclastic +amylocoagulase +amylodextrin +amylodyspepsia +amylogen +amylogenesis +amylogenic +amylohydrolysis +amylohydrolytic +amyloid +amyloidal +amyloidosis +amyloleucite +amylolysis +amylolytic +amylom +amylometer +amylon +amylopectin +amylophagia +amylophosphate +amylophosphoric +amyloplast +amyloplastic +amyloplastid +amylopsin +amylose +amylosis +amylosynthesis +amylum +amyluria +Amynodon +amynodont +amyosthenia +amyosthenic +amyotaxia +amyotonia +amyotrophia +amyotrophic +amyotrophy +amyous +Amyraldism +Amyraldist +Amyridaceae +amyrin +Amyris +amyrol +amyroot +Amytal +amyxorrhea +amyxorrhoea +an +Ana +ana +Anabaena +Anabantidae +Anabaptism +Anabaptist +Anabaptistic +Anabaptistical +Anabaptistically +Anabaptistry +anabaptize +Anabas +anabasine +anabasis +anabasse +anabata +anabathmos +anabatic +anaberoga +anabibazon +anabiosis +anabiotic +Anablepidae +Anableps +anabo +anabohitsite +anabolic +anabolin +anabolism +anabolite +anabolize +anabong +anabranch +anabrosis +anabrotic +anacahuita +anacahuite +anacalypsis +anacampsis +anacamptic +anacamptically +anacamptics +anacamptometer +anacanth +anacanthine +Anacanthini +anacanthous +anacara +anacard +Anacardiaceae +anacardiaceous +anacardic +Anacardium +anacatadidymus +anacatharsis +anacathartic +anacephalaeosis +anacephalize +Anaces +Anacharis +anachorism +anachromasis +anachronic +anachronical +anachronically +anachronism +anachronismatical +anachronist +anachronistic +anachronistical +anachronistically +anachronize +anachronous +anachronously +anachueta +anacid +anacidity +anaclasis +anaclastic +anaclastics +Anaclete +anacleticum +anaclinal +anaclisis +anaclitic +anacoenosis +anacoluthia +anacoluthic +anacoluthically +anacoluthon +anaconda +Anacreon +Anacreontic +Anacreontically +anacrisis +Anacrogynae +anacrogynae +anacrogynous +anacromyodian +anacrotic +anacrotism +anacrusis +anacrustic +anacrustically +anaculture +anacusia +anacusic +anacusis +Anacyclus +anadem +anadenia +anadicrotic +anadicrotism +anadidymus +anadiplosis +anadipsia +anadipsic +anadrom +anadromous +Anadyomene +anaematosis +anaemia +anaemic +anaeretic +anaerobation +anaerobe +anaerobia +anaerobian +anaerobic +anaerobically +anaerobies +anaerobion +anaerobiont +anaerobiosis +anaerobiotic +anaerobiotically +anaerobious +anaerobism +anaerobium +anaerophyte +anaeroplastic +anaeroplasty +anaesthesia +anaesthesiant +anaesthetically +anaesthetizer +anaetiological +anagalactic +Anagallis +anagap +anagenesis +anagenetic +anagep +anagignoskomena +anaglyph +anaglyphic +anaglyphical +anaglyphics +anaglyphoscope +anaglyphy +anaglyptic +anaglyptical +anaglyptics +anaglyptograph +anaglyptographic +anaglyptography +anaglypton +anagnorisis +anagnost +anagoge +anagogic +anagogical +anagogically +anagogics +anagogy +anagram +anagrammatic +anagrammatical +anagrammatically +anagrammatism +anagrammatist +anagrammatize +anagrams +anagraph +anagua +anagyrin +anagyrine +Anagyris +anahau +Anahita +Anaitis +Anakes +anakinesis +anakinetic +anakinetomer +anakinetomeric +anakoluthia +anakrousis +anaktoron +anal +analabos +analav +analcime +analcimite +analcite +analcitite +analecta +analectic +analects +analemma +analemmatic +analepsis +analepsy +analeptic +analeptical +analgen +analgesia +analgesic +Analgesidae +analgesis +analgesist +analgetic +analgia +analgic +analgize +analkalinity +anallagmatic +anallantoic +Anallantoidea +anallantoidean +anallergic +anally +analogic +analogical +analogically +analogicalness +analogion +analogism +analogist +analogistic +analogize +analogon +analogous +analogously +analogousness +analogue +analogy +analphabet +analphabete +analphabetic +analphabetical +analphabetism +analysability +analysable +analysand +analysation +analyse +analyser +analyses +analysis +analyst +analytic +analytical +analytically +analytics +analyzability +analyzable +analyzation +analyze +analyzer +Anam +anam +anama +anamesite +anametadromous +Anamirta +anamirtin +Anamite +anamite +anammonid +anammonide +anamnesis +anamnestic +anamnestically +Anamnia +Anamniata +Anamnionata +anamnionic +Anamniota +anamniote +anamniotic +anamorphic +anamorphism +anamorphoscope +anamorphose +anamorphosis +anamorphote +anamorphous +anan +anana +ananaplas +ananaples +ananas +ananda +anandrarious +anandria +anandrous +ananepionic +anangioid +anangular +Ananias +Ananism +Ananite +anankastic +Anansi +Ananta +anantherate +anantherous +ananthous +ananym +anapaest +anapaestic +anapaestical +anapaestically +anapaganize +anapaite +anapanapa +anapeiratic +anaphalantiasis +Anaphalis +anaphase +Anaphe +anaphia +anaphora +anaphoral +anaphoria +anaphoric +anaphorical +anaphrodisia +anaphrodisiac +anaphroditic +anaphroditous +anaphylactic +anaphylactin +anaphylactogen +anaphylactogenic +anaphylactoid +anaphylatoxin +anaphylaxis +anaphyte +anaplasia +anaplasis +anaplasm +Anaplasma +anaplasmosis +anaplastic +anaplasty +anaplerosis +anaplerotic +anapnea +anapneic +anapnoeic +anapnograph +anapnoic +anapnometer +anapodeictic +anapophysial +anapophysis +anapsid +Anapsida +anapsidan +Anapterygota +anapterygote +anapterygotism +anapterygotous +Anaptomorphidae +Anaptomorphus +anaptotic +anaptychus +anaptyctic +anaptyctical +anaptyxis +anaqua +anarcestean +Anarcestes +anarch +anarchal +anarchial +anarchic +anarchical +anarchically +anarchism +anarchist +anarchistic +anarchize +anarchoindividualist +anarchosocialist +anarchosyndicalism +anarchosyndicalist +anarchy +anarcotin +anareta +anaretic +anaretical +anargyros +anarthria +anarthric +anarthropod +Anarthropoda +anarthropodous +anarthrosis +anarthrous +anarthrously +anarthrousness +anartismos +anarya +Anaryan +Anas +Anasa +anasarca +anasarcous +Anasazi +anaschistic +anaseismic +Anasitch +anaspadias +anaspalin +Anaspida +Anaspidacea +Anaspides +anastalsis +anastaltic +Anastasia +Anastasian +anastasimon +anastasimos +anastasis +Anastasius +anastate +anastatic +Anastatica +Anastatus +anastigmat +anastigmatic +anastomose +anastomosis +anastomotic +Anastomus +anastrophe +Anastrophia +Anat +anatase +anatexis +anathema +anathematic +anathematical +anathematically +anathematism +anathematization +anathematize +anathematizer +anatheme +anathemize +Anatherum +Anatidae +anatifa +Anatifae +anatifer +anatiferous +Anatinacea +Anatinae +anatine +anatocism +Anatole +Anatolian +Anatolic +Anatoly +anatomic +anatomical +anatomically +anatomicobiological +anatomicochirurgical +anatomicomedical +anatomicopathologic +anatomicopathological +anatomicophysiologic +anatomicophysiological +anatomicosurgical +anatomism +anatomist +anatomization +anatomize +anatomizer +anatomopathologic +anatomopathological +anatomy +anatopism +anatox +anatoxin +anatreptic +anatripsis +anatripsology +anatriptic +anatron +anatropal +anatropia +anatropous +Anatum +anaudia +anaunter +anaunters +Anax +Anaxagorean +Anaxagorize +anaxial +Anaximandrian +anaxon +anaxone +Anaxonia +anay +anazoturia +anba +anbury +Ancerata +ancestor +ancestorial +ancestorially +ancestral +ancestrally +ancestress +ancestrial +ancestrian +ancestry +Ancha +Anchat +Anchietea +anchietin +anchietine +anchieutectic +anchimonomineral +Anchisaurus +Anchises +Anchistea +Anchistopoda +anchithere +anchitherioid +anchor +anchorable +anchorage +anchorate +anchored +anchorer +anchoress +anchoret +anchoretic +anchoretical +anchoretish +anchoretism +anchorhold +anchorite +anchoritess +anchoritic +anchoritical +anchoritish +anchoritism +anchorless +anchorlike +anchorwise +anchovy +Anchtherium +Anchusa +anchusin +anchusine +anchylose +anchylosis +ancience +anciency +ancient +ancientism +anciently +ancientness +ancientry +ancienty +ancile +ancilla +ancillary +ancipital +ancipitous +Ancistrocladaceae +ancistrocladaceous +Ancistrocladus +ancistroid +ancon +Ancona +anconad +anconagra +anconal +ancone +anconeal +anconeous +anconeus +anconitis +anconoid +ancony +ancora +ancoral +Ancyloceras +Ancylocladus +Ancylodactyla +ancylopod +Ancylopoda +Ancylostoma +ancylostome +ancylostomiasis +Ancylostomum +Ancylus +Ancyrean +Ancyrene +and +anda +andabatarian +Andalusian +andalusite +Andaman +Andamanese +andante +andantino +Andaqui +Andaquian +Andarko +Andaste +Ande +Andean +Anderson +Andesic +andesine +andesinite +andesite +andesitic +Andevo +Andhra +Andi +Andian +Andine +Andira +andirin +andirine +andiroba +andiron +Andoke +andorite +Andorobo +Andorran +andouillet +andradite +andranatomy +andrarchy +Andre +Andrea +Andreaea +Andreaeaceae +Andreaeales +Andreas +Andrena +andrenid +Andrenidae +Andrew +andrewsite +Andria +Andriana +Andrias +andric +Andries +androcentric +androcephalous +androcephalum +androclinium +Androclus +androconium +androcracy +androcratic +androcyte +androdioecious +androdioecism +androdynamous +androecial +androecium +androgametangium +androgametophore +androgen +androgenesis +androgenetic +androgenic +androgenous +androginous +androgone +androgonia +androgonial +androgonidium +androgonium +Andrographis +andrographolide +androgynal +androgynary +androgyne +androgyneity +androgynia +androgynism +androgynous +androgynus +androgyny +android +androidal +androkinin +androl +androlepsia +androlepsy +Andromache +andromania +Andromaque +Andromeda +Andromede +andromedotoxin +andromonoecious +andromonoecism +andromorphous +andron +Andronicus +andronitis +andropetalar +andropetalous +androphagous +androphobia +androphonomania +androphore +androphorous +androphorum +androphyll +Andropogon +Androsace +Androscoggin +androseme +androsin +androsphinx +androsporangium +androspore +androsterone +androtauric +androtomy +Andy +anear +aneath +anecdota +anecdotage +anecdotal +anecdotalism +anecdote +anecdotic +anecdotical +anecdotically +anecdotist +anele +anelectric +anelectrode +anelectrotonic +anelectrotonus +anelytrous +anematosis +Anemia +anemia +anemic +anemobiagraph +anemochord +anemoclastic +anemogram +anemograph +anemographic +anemographically +anemography +anemological +anemology +anemometer +anemometric +anemometrical +anemometrically +anemometrograph +anemometrographic +anemometrographically +anemometry +anemonal +anemone +Anemonella +anemonin +anemonol +anemony +anemopathy +anemophile +anemophilous +anemophily +Anemopsis +anemoscope +anemosis +anemotaxis +anemotropic +anemotropism +anencephalia +anencephalic +anencephalotrophia +anencephalous +anencephalus +anencephaly +anend +anenergia +anenst +anent +anenterous +anepia +anepigraphic +anepigraphous +anepiploic +anepithymia +anerethisia +aneretic +anergia +anergic +anergy +anerly +aneroid +aneroidograph +anerotic +anerythroplasia +anerythroplastic +anes +anesis +anesthesia +anesthesiant +anesthesimeter +anesthesiologist +anesthesiology +anesthesis +anesthetic +anesthetically +anesthetist +anesthetization +anesthetize +anesthetizer +anesthyl +anethole +Anethum +anetiological +aneuploid +aneuploidy +aneuria +aneuric +aneurilemmic +aneurin +aneurism +aneurismally +aneurysm +aneurysmal +aneurysmally +aneurysmatic +anew +Anezeh +anfractuose +anfractuosity +anfractuous +anfractuousness +anfracture +Angami +Angara +angaralite +angaria +angary +Angdistis +angekok +angel +Angela +angelate +angeldom +Angeleno +angelet +angeleyes +angelfish +angelhood +angelic +Angelica +angelica +Angelical +angelical +angelically +angelicalness +Angelican +angelicic +angelicize +angelico +angelin +Angelina +angeline +angelique +angelize +angellike +Angelo +angelocracy +angelographer +angelolater +angelolatry +angelologic +angelological +angelology +angelomachy +Angelonia +angelophany +angelot +angelship +Angelus +anger +angerly +Angerona +Angeronalia +Angers +Angetenar +Angevin +angeyok +angiasthenia +angico +Angie +angiectasis +angiectopia +angiemphraxis +angiitis +angild +angili +angina +anginal +anginiform +anginoid +anginose +anginous +angioasthenia +angioataxia +angioblast +angioblastic +angiocarditis +angiocarp +angiocarpian +angiocarpic +angiocarpous +angiocavernous +angiocholecystitis +angiocholitis +angiochondroma +angioclast +angiocyst +angiodermatitis +angiodiascopy +angioelephantiasis +angiofibroma +angiogenesis +angiogenic +angiogeny +angioglioma +angiograph +angiography +angiohyalinosis +angiohydrotomy +angiohypertonia +angiohypotonia +angioid +angiokeratoma +angiokinesis +angiokinetic +angioleucitis +angiolipoma +angiolith +angiology +angiolymphitis +angiolymphoma +angioma +angiomalacia +angiomatosis +angiomatous +angiomegaly +angiometer +angiomyocardiac +angiomyoma +angiomyosarcoma +angioneoplasm +angioneurosis +angioneurotic +angionoma +angionosis +angioparalysis +angioparalytic +angioparesis +angiopathy +angiophorous +angioplany +angioplasty +angioplerosis +angiopoietic +angiopressure +angiorrhagia +angiorrhaphy +angiorrhea +angiorrhexis +angiosarcoma +angiosclerosis +angiosclerotic +angioscope +angiosis +angiospasm +angiospastic +angiosperm +Angiospermae +angiospermal +angiospermatous +angiospermic +angiospermous +angiosporous +angiostegnosis +angiostenosis +angiosteosis +angiostomize +angiostomy +angiostrophy +angiosymphysis +angiotasis +angiotelectasia +angiothlipsis +angiotome +angiotomy +angiotonic +angiotonin +angiotribe +angiotripsy +angiotrophic +Angka +anglaise +angle +angleberry +angled +anglehook +anglepod +angler +Angles +anglesite +anglesmith +angletouch +angletwitch +anglewing +anglewise +angleworm +Anglian +Anglic +Anglican +Anglicanism +Anglicanize +Anglicanly +Anglicanum +Anglicism +Anglicist +Anglicization +anglicization +Anglicize +anglicize +Anglification +Anglify +anglimaniac +angling +Anglish +Anglist +Anglistics +Anglogaea +Anglogaean +angloid +Angloman +Anglomane +Anglomania +Anglomaniac +Anglophile +Anglophobe +Anglophobia +Anglophobiac +Anglophobic +Anglophobist +ango +Angola +angolar +Angolese +angor +Angora +angostura +Angouleme +Angoumian +Angraecum +angrily +angriness +angrite +angry +angst +angster +Angstrom +angstrom +anguid +Anguidae +anguiform +Anguilla +Anguillaria +Anguillidae +anguilliform +anguilloid +Anguillula +Anguillulidae +Anguimorpha +anguine +anguineal +anguineous +Anguinidae +anguiped +Anguis +anguis +anguish +anguished +anguishful +anguishous +anguishously +angula +angular +angulare +angularity +angularization +angularize +angularly +angularness +angulate +angulated +angulately +angulateness +angulation +angulatogibbous +angulatosinuous +anguliferous +angulinerved +Anguloa +angulodentate +angulometer +angulosity +angulosplenial +angulous +anguria +Angus +angusticlave +angustifoliate +angustifolious +angustirostrate +angustisellate +angustiseptal +angustiseptate +angwantibo +anhalamine +anhaline +anhalonine +Anhalonium +anhalouidine +anhang +Anhanga +anharmonic +anhedonia +anhedral +anhedron +anhelation +anhelous +anhematosis +anhemolytic +anhidrosis +anhidrotic +anhima +Anhimae +Anhimidae +anhinga +anhistic +anhistous +anhungered +anhungry +anhydrate +anhydration +anhydremia +anhydremic +anhydric +anhydride +anhydridization +anhydridize +anhydrite +anhydrization +anhydrize +anhydroglocose +anhydromyelia +anhydrous +anhydroxime +anhysteretic +ani +Aniba +Anice +aniconic +aniconism +anicular +anicut +anidian +anidiomatic +anidiomatical +anidrosis +Aniellidae +aniente +anigh +anight +anights +anil +anilao +anilau +anile +anileness +anilic +anilid +anilide +anilidic +anilidoxime +aniline +anilinism +anilinophile +anilinophilous +anility +anilla +anilopyrin +anilopyrine +anima +animability +animable +animableness +animadversion +animadversional +animadversive +animadversiveness +animadvert +animadverter +animal +animalcula +animalculae +animalcular +animalcule +animalculine +animalculism +animalculist +animalculous +animalculum +animalhood +Animalia +animalian +animalic +animalier +animalish +animalism +animalist +animalistic +animality +Animalivora +animalivore +animalivorous +animalization +animalize +animally +animastic +animastical +animate +animated +animatedly +animately +animateness +animater +animating +animatingly +animation +animatism +animatistic +animative +animatograph +animator +anime +animi +Animikean +animikite +animism +animist +animistic +animize +animosity +animotheism +animous +animus +anion +anionic +aniridia +anis +anisal +anisalcohol +anisaldehyde +anisaldoxime +anisamide +anisandrous +anisanilide +anisate +anischuria +anise +aniseed +aniseikonia +aniseikonic +aniselike +aniseroot +anisette +anisic +anisidin +anisidine +anisil +anisilic +anisobranchiate +anisocarpic +anisocarpous +anisocercal +anisochromatic +anisochromia +anisocoria +anisocotyledonous +anisocotyly +anisocratic +anisocycle +anisocytosis +anisodactyl +Anisodactyla +Anisodactyli +anisodactylic +anisodactylous +anisodont +anisogamete +anisogamous +anisogamy +anisogenous +anisogeny +anisognathism +anisognathous +anisogynous +anisoin +anisole +anisoleucocytosis +Anisomeles +anisomelia +anisomelus +anisomeric +anisomerous +anisometric +anisometrope +anisometropia +anisometropic +anisomyarian +Anisomyodi +anisomyodian +anisomyodous +anisopetalous +anisophyllous +anisophylly +anisopia +anisopleural +anisopleurous +anisopod +Anisopoda +anisopodal +anisopodous +anisopogonous +Anisoptera +anisopterous +anisosepalous +anisospore +anisostaminous +anisostemonous +anisosthenic +anisostichous +Anisostichus +anisostomous +anisotonic +anisotropal +anisotrope +anisotropic +anisotropical +anisotropically +anisotropism +anisotropous +anisotropy +anisoyl +anisum +anisuria +anisyl +anisylidene +Anita +anither +anitrogenous +anjan +Anjou +ankaramite +ankaratrite +ankee +anker +ankerite +ankh +ankle +anklebone +anklejack +anklet +anklong +Ankoli +Ankou +ankus +ankusha +ankylenteron +ankyloblepharon +ankylocheilia +ankylodactylia +ankylodontia +ankyloglossia +ankylomele +ankylomerism +ankylophobia +ankylopodia +ankylopoietic +ankyloproctia +ankylorrhinia +Ankylosaurus +ankylose +ankylosis +ankylostoma +ankylotia +ankylotic +ankylotome +ankylotomy +ankylurethria +ankyroid +anlace +anlaut +Ann +ann +Anna +anna +Annabel +annabergite +annal +annale +annaline +annalism +annalist +annalistic +annalize +annals +Annam +Annamese +Annamite +Annamitic +Annapurna +Annard +annat +annates +annatto +Anne +anneal +annealer +annectent +annection +annelid +Annelida +annelidan +Annelides +annelidian +annelidous +annelism +Annellata +anneloid +annerodite +Anneslia +annet +Annette +annex +annexa +annexable +annexal +annexation +annexational +annexationist +annexer +annexion +annexionist +annexitis +annexive +annexment +annexure +annidalin +Annie +Anniellidae +annihilability +annihilable +annihilate +annihilation +annihilationism +annihilationist +annihilative +annihilator +annihilatory +Annist +annite +anniversarily +anniversariness +anniversary +anniverse +annodated +Annona +annona +Annonaceae +annonaceous +annotate +annotater +annotation +annotative +annotator +annotatory +annotine +annotinous +announce +announceable +announcement +announcer +annoy +annoyance +annoyancer +annoyer +annoyful +annoying +annoyingly +annoyingness +annoyment +annual +annualist +annualize +annually +annuary +annueler +annuent +annuitant +annuity +annul +annular +Annularia +annularity +annularly +annulary +Annulata +annulate +annulated +annulation +annulet +annulettee +annulism +annullable +annullate +annullation +annuller +annulment +annuloid +Annuloida +Annulosa +annulosan +annulose +annulus +annunciable +annunciate +annunciation +annunciative +annunciator +annunciatory +anoa +Anobiidae +anocarpous +anociassociation +anococcygeal +anodal +anode +anodendron +anodic +anodically +anodize +Anodon +Anodonta +anodontia +anodos +anodyne +anodynia +anodynic +anodynous +anoegenetic +anoesia +anoesis +anoestrous +anoestrum +anoestrus +anoetic +anogenic +anogenital +Anogra +anoil +anoine +anoint +anointer +anointment +anole +anoli +anolian +Anolis +Anolympiad +anolyte +Anomala +anomaliflorous +anomaliped +anomalism +anomalist +anomalistic +anomalistical +anomalistically +anomalocephalus +anomaloflorous +Anomalogonatae +anomalogonatous +Anomalon +anomalonomy +Anomalopteryx +anomaloscope +anomalotrophy +anomalous +anomalously +anomalousness +anomalure +Anomaluridae +Anomalurus +anomaly +Anomatheca +Anomia +Anomiacea +Anomiidae +anomite +anomocarpous +anomodont +Anomodontia +Anomoean +Anomoeanism +anomophyllous +anomorhomboid +anomorhomboidal +anomphalous +Anomura +anomural +anomuran +anomurous +anomy +anon +anonang +anoncillo +anonol +anonychia +anonym +anonyma +anonymity +anonymous +anonymously +anonymousness +anonymuncule +anoopsia +anoperineal +anophele +Anopheles +Anophelinae +anopheline +anophoria +anophthalmia +anophthalmos +Anophthalmus +anophyte +anopia +anopisthographic +Anopla +Anoplanthus +anoplocephalic +anoplonemertean +Anoplonemertini +anoplothere +Anoplotheriidae +anoplotherioid +Anoplotherium +anoplotheroid +Anoplura +anopluriform +anopsia +anopubic +anorak +anorchia +anorchism +anorchous +anorchus +anorectal +anorectic +anorectous +anorexia +anorexy +anorgana +anorganic +anorganism +anorganology +anormal +anormality +anorogenic +anorth +anorthic +anorthite +anorthitic +anorthitite +anorthoclase +anorthographic +anorthographical +anorthographically +anorthography +anorthophyre +anorthopia +anorthoscope +anorthose +anorthosite +anoscope +anoscopy +Anosia +anosmatic +anosmia +anosmic +anosphrasia +anosphresia +anospinal +anostosis +Anostraca +anoterite +another +anotherkins +anotia +anotropia +anotta +anotto +anotus +anounou +Anous +anovesical +anoxemia +anoxemic +anoxia +anoxic +anoxidative +anoxybiosis +anoxybiotic +anoxyscope +ansa +ansar +ansarian +Ansarie +ansate +ansation +Anseis +Ansel +Anselm +Anselmian +Anser +anserated +Anseres +Anseriformes +Anserinae +anserine +anserous +anspessade +ansu +ansulate +answer +answerability +answerable +answerableness +answerably +answerer +answeringly +answerless +answerlessly +ant +Anta +anta +antacid +antacrid +antadiform +Antaean +Antaeus +antagonism +antagonist +antagonistic +antagonistical +antagonistically +antagonization +antagonize +antagonizer +antagony +Antaimerina +Antaios +Antaiva +antal +antalgesic +antalgol +antalkali +antalkaline +antambulacral +antanacathartic +antanaclasis +Antanandro +antanemic +antapex +antaphrodisiac +antaphroditic +antapocha +antapodosis +antapology +antapoplectic +Antar +Antara +antarchism +antarchist +antarchistic +antarchistical +antarchy +Antarctalia +Antarctalian +antarctic +Antarctica +antarctica +antarctical +antarctically +Antarctogaea +Antarctogaean +Antares +antarthritic +antasphyctic +antasthenic +antasthmatic +antatrophic +antdom +ante +anteact +anteal +anteambulate +anteambulation +anteater +antebaptismal +antebath +antebrachial +antebrachium +antebridal +antecabinet +antecaecal +antecardium +antecavern +antecedaneous +antecedaneously +antecede +antecedence +antecedency +antecedent +antecedental +antecedently +antecessor +antechamber +antechapel +Antechinomys +antechoir +antechurch +anteclassical +antecloset +antecolic +antecommunion +anteconsonantal +antecornu +antecourt +antecoxal +antecubital +antecurvature +antedate +antedawn +antediluvial +antediluvially +antediluvian +Antedon +antedonin +antedorsal +antefebrile +antefix +antefixal +anteflected +anteflexed +anteflexion +antefurca +antefurcal +antefuture +antegarden +antegrade +antehall +antehistoric +antehuman +antehypophysis +anteinitial +antejentacular +antejudiciary +antejuramentum +antelabium +antelegal +antelocation +antelope +antelopian +antelucan +antelude +anteluminary +antemarginal +antemarital +antemedial +antemeridian +antemetallic +antemetic +antemillennial +antemingent +antemortal +antemundane +antemural +antenarial +antenatal +antenatalitial +antenati +antenave +antenna +antennae +antennal +Antennaria +antennariid +Antennariidae +Antennarius +antennary +Antennata +antennate +antenniferous +antenniform +antennula +antennular +antennulary +antennule +antenodal +antenoon +Antenor +antenumber +anteoccupation +anteocular +anteopercle +anteoperculum +anteorbital +antepagmenta +antepagments +antepalatal +antepaschal +antepast +antepatriarchal +antepectoral +antepectus +antependium +antepenult +antepenultima +antepenultimate +antephialtic +antepileptic +antepirrhema +anteporch +anteportico +anteposition +anteposthumous +anteprandial +antepredicament +antepredicamental +antepreterit +antepretonic +anteprohibition +anteprostate +anteprostatic +antepyretic +antequalm +antereformation +antereformational +anteresurrection +anterethic +anterevolutional +anterevolutionary +anteriad +anterior +anteriority +anteriorly +anteriorness +anteroclusion +anterodorsal +anteroexternal +anterofixation +anteroflexion +anterofrontal +anterograde +anteroinferior +anterointerior +anterointernal +anterolateral +anterolaterally +anteromedial +anteromedian +anteroom +anteroparietal +anteroposterior +anteroposteriorly +anteropygal +anterospinal +anterosuperior +anteroventral +anteroventrally +antes +antescript +antesignanus +antespring +antestature +antesternal +antesternum +antesunrise +antesuperior +antetemple +antetype +Anteva +antevenient +anteversion +antevert +antevocalic +antewar +anthecological +anthecologist +anthecology +Antheia +anthela +anthelion +anthelmintic +anthem +anthema +anthemene +anthemia +Anthemideae +anthemion +Anthemis +anthemwise +anthemy +anther +Antheraea +antheral +Anthericum +antherid +antheridial +antheridiophore +antheridium +antheriferous +antheriform +antherless +antherogenous +antheroid +antherozoid +antherozoidal +antherozooid +antherozooidal +anthesis +Anthesteria +Anthesteriac +anthesterin +Anthesterion +anthesterol +antheximeter +Anthicidae +Anthidium +anthill +Anthinae +anthine +anthobiology +anthocarp +anthocarpous +anthocephalous +Anthoceros +Anthocerotaceae +Anthocerotales +anthocerote +anthochlor +anthochlorine +anthoclinium +anthocyan +anthocyanidin +anthocyanin +anthodium +anthoecological +anthoecologist +anthoecology +anthogenesis +anthogenetic +anthogenous +anthography +anthoid +anthokyan +antholite +anthological +anthologically +anthologion +anthologist +anthologize +anthology +antholysis +Antholyza +anthomania +anthomaniac +Anthomedusae +anthomedusan +Anthomyia +anthomyiid +Anthomyiidae +Anthonin +Anthonomus +Anthony +anthood +anthophagous +Anthophila +anthophile +anthophilian +anthophilous +anthophobia +Anthophora +anthophore +Anthophoridae +anthophorous +anthophyllite +anthophyllitic +Anthophyta +anthophyte +anthorine +anthosiderite +Anthospermum +anthotaxis +anthotaxy +anthotropic +anthotropism +anthoxanthin +Anthoxanthum +Anthozoa +anthozoan +anthozoic +anthozooid +anthozoon +anthracemia +anthracene +anthraceniferous +anthrachrysone +anthracia +anthracic +anthraciferous +anthracin +anthracite +anthracitic +anthracitiferous +anthracitious +anthracitism +anthracitization +anthracnose +anthracnosis +anthracocide +anthracoid +anthracolithic +anthracomancy +Anthracomarti +anthracomartian +Anthracomartus +anthracometer +anthracometric +anthraconecrosis +anthraconite +Anthracosaurus +anthracosis +anthracothere +Anthracotheriidae +Anthracotherium +anthracotic +anthracyl +anthradiol +anthradiquinone +anthraflavic +anthragallol +anthrahydroquinone +anthramine +anthranil +anthranilate +anthranilic +anthranol +anthranone +anthranoyl +anthranyl +anthraphenone +anthrapurpurin +anthrapyridine +anthraquinol +anthraquinone +anthraquinonyl +anthrarufin +anthratetrol +anthrathiophene +anthratriol +anthrax +anthraxolite +anthraxylon +Anthrenus +anthribid +Anthribidae +Anthriscus +anthrohopobiological +anthroic +anthrol +anthrone +anthropic +anthropical +Anthropidae +anthropobiologist +anthropobiology +anthropocentric +anthropocentrism +anthropoclimatologist +anthropoclimatology +anthropocosmic +anthropodeoxycholic +Anthropodus +anthropogenesis +anthropogenetic +anthropogenic +anthropogenist +anthropogenous +anthropogeny +anthropogeographer +anthropogeographical +anthropogeography +anthropoglot +anthropogony +anthropography +anthropoid +anthropoidal +Anthropoidea +anthropoidean +anthropolater +anthropolatric +anthropolatry +anthropolite +anthropolithic +anthropolitic +anthropological +anthropologically +anthropologist +anthropology +anthropomancy +anthropomantic +anthropomantist +anthropometer +anthropometric +anthropometrical +anthropometrically +anthropometrist +anthropometry +anthropomorph +Anthropomorpha +anthropomorphic +anthropomorphical +anthropomorphically +Anthropomorphidae +anthropomorphism +anthropomorphist +anthropomorphite +anthropomorphitic +anthropomorphitical +anthropomorphitism +anthropomorphization +anthropomorphize +anthropomorphological +anthropomorphologically +anthropomorphology +anthropomorphosis +anthropomorphotheist +anthropomorphous +anthropomorphously +anthroponomical +anthroponomics +anthroponomist +anthroponomy +anthropopathia +anthropopathic +anthropopathically +anthropopathism +anthropopathite +anthropopathy +anthropophagi +anthropophagic +anthropophagical +anthropophaginian +anthropophagism +anthropophagist +anthropophagistic +anthropophagite +anthropophagize +anthropophagous +anthropophagously +anthropophagy +anthropophilous +anthropophobia +anthropophuism +anthropophuistic +anthropophysiography +anthropophysite +Anthropopithecus +anthropopsychic +anthropopsychism +Anthropos +anthroposcopy +anthroposociologist +anthroposociology +anthroposomatology +anthroposophical +anthroposophist +anthroposophy +anthropoteleoclogy +anthropoteleological +anthropotheism +anthropotomical +anthropotomist +anthropotomy +anthropotoxin +Anthropozoic +anthropurgic +anthroropolith +anthroxan +anthroxanic +anthryl +anthrylene +Anthurium +Anthus +Anthyllis +anthypophora +anthypophoretic +Anti +anti +antiabolitionist +antiabrasion +antiabrin +antiabsolutist +antiacid +antiadiaphorist +antiaditis +antiadministration +antiae +antiaesthetic +antiager +antiagglutinating +antiagglutinin +antiaggression +antiaggressionist +antiaggressive +antiaircraft +antialbumid +antialbumin +antialbumose +antialcoholic +antialcoholism +antialcoholist +antialdoxime +antialexin +antialien +antiamboceptor +antiamusement +antiamylase +antianaphylactogen +antianaphylaxis +antianarchic +antianarchist +antiangular +antiannexation +antiannexationist +antianopheline +antianthrax +antianthropocentric +antianthropomorphism +antiantibody +antiantidote +antiantienzyme +antiantitoxin +antiaphrodisiac +antiaphthic +antiapoplectic +antiapostle +antiaquatic +antiar +Antiarcha +Antiarchi +antiarin +Antiaris +antiaristocrat +antiarthritic +antiascetic +antiasthmatic +antiastronomical +antiatheism +antiatheist +antiatonement +antiattrition +antiautolysin +antibacchic +antibacchius +antibacterial +antibacteriolytic +antiballooner +antibalm +antibank +antibasilican +antibenzaldoxime +antiberiberin +antibibliolatry +antibigotry +antibilious +antibiont +antibiosis +antibiotic +antibishop +antiblastic +antiblennorrhagic +antiblock +antiblue +antibody +antiboxing +antibreakage +antibridal +antibromic +antibubonic +Antiburgher +antic +anticachectic +antical +anticalcimine +anticalculous +anticalligraphic +anticancer +anticapital +anticapitalism +anticapitalist +anticardiac +anticardium +anticarious +anticarnivorous +anticaste +anticatalase +anticatalyst +anticatalytic +anticatalyzer +anticatarrhal +anticathexis +anticathode +anticaustic +anticensorship +anticentralization +anticephalalgic +anticeremonial +anticeremonialism +anticeremonialist +anticheater +antichlor +antichlorine +antichloristic +antichlorotic +anticholagogue +anticholinergic +antichoromanic +antichorus +antichresis +antichretic +antichrist +antichristian +antichristianity +antichristianly +antichrome +antichronical +antichronically +antichthon +antichurch +antichurchian +antichymosin +anticipant +anticipatable +anticipate +anticipation +anticipative +anticipatively +anticipator +anticipatorily +anticipatory +anticivic +anticivism +anticize +anticker +anticlactic +anticlassical +anticlassicist +Anticlea +anticlergy +anticlerical +anticlericalism +anticlimactic +anticlimax +anticlinal +anticline +anticlinorium +anticlockwise +anticlogging +anticly +anticnemion +anticness +anticoagulant +anticoagulating +anticoagulative +anticoagulin +anticogitative +anticolic +anticombination +anticomet +anticomment +anticommercial +anticommunist +anticomplement +anticomplementary +anticomplex +anticonceptionist +anticonductor +anticonfederationist +anticonformist +anticonscience diff --git a/tests/bench/fmtFlatten.lean b/tests/bench/fmtFlatten.lean new file mode 100644 index 000000000000..fc7704c43a99 --- /dev/null +++ b/tests/bench/fmtFlatten.lean @@ -0,0 +1,28 @@ +import Lean.Data.Fmt.Formatter + +open Lean.Fmt + +def quadratic (n : Nat) : Doc := + if n = 0 then + .text "line" + else + .maybeFlattened + (Doc.joinUsing .nl #[quadratic (n - 1), .text "line"]) + +@[noinline] +def doc (n : Nat) : IO Doc := + return quadratic n + +@[noinline] +def format (doc : Doc) : IO (Option String) := do + return format? doc 80 100 + +def main (args : List String) : IO Unit := do + let n := (args[0]!).toNat! + let d ← doc n + let startNs ← IO.monoNanosNow + let r? ← format d + let endNs ← IO.monoNanosNow + let benchTime : Float := (endNs - startNs).toFloat / 1_000_000_000.0 + assert! r?.isSome + IO.println s!"format: {benchTime}" diff --git a/tests/bench/fmtJson.lean b/tests/bench/fmtJson.lean new file mode 100644 index 000000000000..5bb9fa0d72c8 --- /dev/null +++ b/tests/bench/fmtJson.lean @@ -0,0 +1,62 @@ +import Lean.Data.Fmt.Formatter +import Lean.Data.Json + +open Lean.Fmt + +def hConcat (ds : List Doc) : Doc := + match ds with + | [] => .failure + | [d] => d + | d :: ds => + ds.foldl (init := d) fun acc d => (Doc.flattened acc).append d + +def encloseSep (left right sep : Doc) (ds : List Doc) : Doc := + match ds with + | [] => .append left (.aligned right) + | [d] => .join #[left, .aligned d, .aligned right] + | d :: ds => + .append + (.either + (hConcat (left :: (d :: ds).intersperse sep)) + (Doc.joinUsing .hardNl (Doc.append left (.aligned d) :: ds.map (fun d => Doc.append sep (.aligned d))).toArray)) + (.aligned right) + +partial def pp (j : Lean.Json) : Doc := + match j with + | .null => .text "null" + | .bool false => .text "false" + | .bool true => .text "true" + | .num n => .text n.toString + | .str s => .text s!"\"{s}\"" + | .arr a => + let a := a.map pp + encloseSep (.text "[") (.text "]") (.text ",") a.toList + | .obj kvPairs => + let kvPairs := kvPairs.toList.map fun (k, v) => + let k := .text s!"\"{k}\": " + let v := pp v + Doc.append k (.aligned v) + encloseSep (.text "{") (.text "}") (.text ",") kvPairs + +def readJson : IO Lean.Json := do + let c ← IO.FS.readFile "./lean/run/fmtJson1k.json" + IO.ofExcept <| Lean.Json.parse c + +@[noinline] +def doc : IO Doc := do + let json ← IO.FS.readFile "fmtJson10k.json" + let json ← IO.ofExcept <| Lean.Json.parse json + return pp json + +@[noinline] +def format (doc : Doc) : IO (Option String) := do + return format? doc 80 100 + +def main (_ : List String) : IO Unit := do + let d ← doc + let startNs ← IO.monoNanosNow + let r? ← format d + let endNs ← IO.monoNanosNow + let benchTime : Float := (endNs - startNs).toFloat / 1_000_000_000.0 + assert! r?.isSome + IO.println s!"format: {benchTime}" diff --git a/tests/bench/fmtJson10k.json b/tests/bench/fmtJson10k.json new file mode 100644 index 000000000000..818a64c22428 --- /dev/null +++ b/tests/bench/fmtJson10k.json @@ -0,0 +1,10562 @@ +[ + { + "_id": "58fa646dce2ac8904cb09b67", + "index": 0, + "guid": "fe99b725-30e8-4706-bd4d-c2b6f1a526e3", + "isActive": true, + "balance": "$1,134.56", + "picture": "http://placehold.it/32x32", + "age": 29, + "eyeColor": "green", + "name": "Flossie Stout", + "gender": "female", + "company": "SENMEI", + "email": "flossiestout@senmei.com", + "phone": "+1 (946) 540-2215", + "address": "486 Pilling Street, Hollymead, Washington, 6043", + "registered": "2014-03-24T04:12:59 -01:00", + "latitude": -24.212991, + "longitude": -170.376913, + "tags": [ + "fugiat", + "ullamco", + "voluptate", + "cillum", + "enim", + "ut", + "laboris" + ], + "friends": [ + { + "id": 0, + "name": "Roslyn Miranda" + }, + { + "id": 1, + "name": "Lolita Kidd" + }, + { + "id": 2, + "name": "Amparo Russell" + } + ], + "greeting": "Hello, Flossie Stout! You have 1 unread messages.", + "favoriteFruit": "apple" + }, + { + "_id": "58fa646da7e2491b6e91074a", + "index": 1, + "guid": "1d5e6f80-da44-4c4f-8b03-d69031d0b3c8", + "isActive": true, + "balance": "$1,948.61", + "picture": "http://placehold.it/32x32", + "age": 30, + "eyeColor": "blue", + "name": "Avis Madden", + "gender": "female", + "company": "CINASTER", + "email": "avismadden@cinaster.com", + "phone": "+1 (981) 549-3574", + "address": "544 Grattan Street, Bodega, Colorado, 7270", + "registered": "2017-01-22T11:21:35 -01:00", + "latitude": 39.720414, + "longitude": 36.95094, + "tags": [ + "ad", + "officia", + "aliquip", + "dolore", + "fugiat", + "officia", + "et" + ], + "friends": [ + { + "id": 0, + "name": "Lucinda Moran" + }, + { + "id": 1, + "name": "Hendrix Mcgee" + }, + { + "id": 2, + "name": "Benton Mcintyre" + } + ], + "greeting": "Hello, Avis Madden! You have 10 unread messages.", + "favoriteFruit": "strawberry" + }, + { + "_id": "58fa646d9cf8de8500c85eb2", + "index": 2, + "guid": "8f182512-f853-46ff-9677-34cbe2159233", + "isActive": false, + "balance": "$3,489.34", + "picture": "http://placehold.it/32x32", + "age": 36, + "eyeColor": "brown", + "name": "Rhea Barnes", + "gender": "female", + "company": "FRANSCENE", + "email": "rheabarnes@franscene.com", + "phone": "+1 (863) 545-3827", + "address": "588 Malta Street, Venice, Puerto Rico, 8786", + "registered": "2017-01-14T12:24:37 -01:00", + "latitude": -22.195949, + "longitude": -147.278613, + "tags": [ + "aliquip", + "proident", + "officia", + "dolore", + "minim", + "aute", + "amet" + ], + "friends": [ + { + "id": 0, + "name": "Yesenia Leblanc" + }, + { + "id": 1, + "name": "Renee Wynn" + }, + { + "id": 2, + "name": "Mendez Newman" + } + ], + "greeting": "Hello, Rhea Barnes! You have 6 unread messages.", + "favoriteFruit": "apple" + }, + { + "_id": "58fa646d6e36e937a4728bb3", + "index": 3, + "guid": "769757e8-7155-4434-80d3-5f1ad10c4e08", + "isActive": false, + "balance": "$3,719.81", + "picture": "http://placehold.it/32x32", + "age": 23, + "eyeColor": "brown", + "name": "Christian Pennington", + "gender": "female", + "company": "CINCYR", + "email": "christianpennington@cincyr.com", + "phone": "+1 (968) 594-2599", + "address": "958 Jerome Avenue, Hiwasse, Rhode Island, 7045", + "registered": "2017-03-15T10:04:32 -01:00", + "latitude": -55.681155, + "longitude": 32.631437, + "tags": [ + "adipisicing", + "veniam", + "deserunt", + "anim", + "reprehenderit", + "velit", + "duis" + ], + "friends": [ + { + "id": 0, + "name": "Noelle Carrillo" + }, + { + "id": 1, + "name": "Conley Hall" + }, + { + "id": 2, + "name": "Harrington Stuart" + } + ], + "greeting": "Hello, Christian Pennington! You have 4 unread messages.", + "favoriteFruit": "apple" + }, + { + "_id": "58fa646d0f4dcff2845edec1", + "index": 4, + "guid": "94dea3e2-8725-4e5f-b507-fac6e18a6d2b", + "isActive": false, + "balance": "$2,037.38", + "picture": "http://placehold.it/32x32", + "age": 32, + "eyeColor": "green", + "name": "Valencia Booker", + "gender": "male", + "company": "MAZUDA", + "email": "valenciabooker@mazuda.com", + "phone": "+1 (909) 451-3109", + "address": "458 Bevy Court, Stouchsburg, Alabama, 1952", + "registered": "2016-04-02T10:39:51 -02:00", + "latitude": 56.87185, + "longitude": -95.763344, + "tags": [ + "tempor", + "ad", + "elit", + "commodo", + "proident", + "et", + "tempor" + ], + "friends": [ + { + "id": 0, + "name": "Nola Delaney" + }, + { + "id": 1, + "name": "Whitehead Howe" + }, + { + "id": 2, + "name": "Luella Knight" + } + ], + "greeting": "Hello, Valencia Booker! You have 8 unread messages.", + "favoriteFruit": "apple" + }, + { + "_id": "58fa646d93e65bb1b3f51032", + "index": 5, + "guid": "de2da198-9093-4638-9ef5-223025084c75", + "isActive": false, + "balance": "$2,737.19", + "picture": "http://placehold.it/32x32", + "age": 23, + "eyeColor": "brown", + "name": "Simon Parrish", + "gender": "male", + "company": "GROK", + "email": "simonparrish@grok.com", + "phone": "+1 (931) 466-3405", + "address": "899 Wallabout Street, Warren, Wisconsin, 1620", + "registered": "2016-05-06T05:58:57 -02:00", + "latitude": 38.608688, + "longitude": 47.125848, + "tags": [ + "cupidatat", + "aute", + "ad", + "minim", + "occaecat", + "cupidatat", + "fugiat" + ], + "friends": [ + { + "id": 0, + "name": "Marie Whitehead" + }, + { + "id": 1, + "name": "Harvey Rogers" + }, + { + "id": 2, + "name": "Mcmahon Keller" + } + ], + "greeting": "Hello, Simon Parrish! You have 7 unread messages.", + "favoriteFruit": "banana" + }, + { + "_id": "58fa646daa17b73db43d4650", + "index": 6, + "guid": "2ce42035-f709-42bb-af8d-8fce008dedf5", + "isActive": true, + "balance": "$3,052.77", + "picture": "http://placehold.it/32x32", + "age": 39, + "eyeColor": "green", + "name": "Salas Hernandez", + "gender": "male", + "company": "INSURITY", + "email": "salashernandez@insurity.com", + "phone": "+1 (800) 487-3315", + "address": "857 Glenmore Avenue, Bowden, California, 2303", + "registered": "2016-05-21T11:40:44 -02:00", + "latitude": -81.381604, + "longitude": -43.534619, + "tags": [ + "Lorem", + "nulla", + "sint", + "enim", + "in", + "aute", + "proident" + ], + "friends": [ + { + "id": 0, + "name": "Hensley Everett" + }, + { + "id": 1, + "name": "Jeanne Buchanan" + }, + { + "id": 2, + "name": "Puckett Hester" + } + ], + "greeting": "Hello, Salas Hernandez! You have 8 unread messages.", + "favoriteFruit": "banana" + }, + { + "_id": "58fa646db7dfc72c7974764b", + "index": 7, + "guid": "9083e721-1221-481d-addd-1040e284944e", + "isActive": true, + "balance": "$2,863.88", + "picture": "http://placehold.it/32x32", + "age": 24, + "eyeColor": "blue", + "name": "Janie Puckett", + "gender": "female", + "company": "HATOLOGY", + "email": "janiepuckett@hatology.com", + "phone": "+1 (889) 590-2440", + "address": "256 Hart Place, Churchill, Arkansas, 4794", + "registered": "2016-12-26T10:58:10 -01:00", + "latitude": -2.85166, + "longitude": -71.677735, + "tags": [ + "excepteur", + "amet", + "elit", + "ut", + "sunt", + "proident", + "non" + ], + "friends": [ + { + "id": 0, + "name": "Letha Morin" + }, + { + "id": 1, + "name": "Bonnie Aguirre" + }, + { + "id": 2, + "name": "Paulette Mendoza" + } + ], + "greeting": "Hello, Janie Puckett! You have 2 unread messages.", + "favoriteFruit": "strawberry" + }, + { + "_id": "58fa646d1b50aa3872da0892", + "index": 8, + "guid": "e0ebcfe5-956f-4a7e-bc8f-5fb90c38c192", + "isActive": true, + "balance": "$3,748.01", + "picture": "http://placehold.it/32x32", + "age": 31, + "eyeColor": "brown", + "name": "Jill Collier", + "gender": "female", + "company": "ZAPPIX", + "email": "jillcollier@zappix.com", + "phone": "+1 (965) 471-3646", + "address": "228 Ridgewood Avenue, Dexter, Vermont, 8708", + "registered": "2014-08-20T05:11:38 -02:00", + "latitude": -31.07983, + "longitude": -118.160665, + "tags": [ + "ea", + "labore", + "exercitation", + "nisi", + "voluptate", + "velit", + "do" + ], + "friends": [ + { + "id": 0, + "name": "Kerri Frost" + }, + { + "id": 1, + "name": "Sylvia Koch" + }, + { + "id": 2, + "name": "Addie Eaton" + } + ], + "greeting": "Hello, Jill Collier! You have 9 unread messages.", + "favoriteFruit": "apple" + }, + { + "_id": "58fa646d42102436b4a27348", + "index": 9, + "guid": "ca68ccb5-2eeb-4e69-97e9-8b9a84f98c3c", + "isActive": true, + "balance": "$2,124.95", + "picture": "http://placehold.it/32x32", + "age": 24, + "eyeColor": "brown", + "name": "Chapman Ayala", + "gender": "male", + "company": "GUSHKOOL", + "email": "chapmanayala@gushkool.com", + "phone": "+1 (977) 595-3236", + "address": "168 Fleet Walk, Moraida, New Hampshire, 705", + "registered": "2014-11-03T10:29:49 -01:00", + "latitude": -53.335234, + "longitude": -36.283822, + "tags": [ + "deserunt", + "non", + "ea", + "sit", + "qui", + "ut", + "est" + ], + "friends": [ + { + "id": 0, + "name": "Rachel Benson" + }, + { + "id": 1, + "name": "Gertrude Becker" + }, + { + "id": 2, + "name": "Pam Decker" + } + ], + "greeting": "Hello, Chapman Ayala! You have 6 unread messages.", + "favoriteFruit": "apple" + }, + { + "_id": "58fa646d75c0e2299af72d31", + "index": 10, + "guid": "2c76a7f4-67a7-4e4a-a221-a22f5860297b", + "isActive": false, + "balance": "$3,142.14", + "picture": "http://placehold.it/32x32", + "age": 35, + "eyeColor": "blue", + "name": "Greta Maxwell", + "gender": "female", + "company": "SURELOGIC", + "email": "gretamaxwell@surelogic.com", + "phone": "+1 (850) 445-2538", + "address": "206 Forrest Street, Sultana, Delaware, 3672", + "registered": "2015-04-12T01:22:16 -02:00", + "latitude": 1.877579, + "longitude": -71.637562, + "tags": [ + "qui", + "reprehenderit", + "veniam", + "ullamco", + "id", + "pariatur", + "do" + ], + "friends": [ + { + "id": 0, + "name": "Kathy Santana" + }, + { + "id": 1, + "name": "Laurie Mcpherson" + }, + { + "id": 2, + "name": "Michele Valencia" + } + ], + "greeting": "Hello, Greta Maxwell! You have 6 unread messages.", + "favoriteFruit": "strawberry" + }, + { + "_id": "58fa646d5fa5cb726044d848", + "index": 11, + "guid": "9f96664f-b870-47f8-9a83-251ce76008ad", + "isActive": false, + "balance": "$2,245.46", + "picture": "http://placehold.it/32x32", + "age": 38, + "eyeColor": "blue", + "name": "Julie Gregory", + "gender": "female", + "company": "MOMENTIA", + "email": "juliegregory@momentia.com", + "phone": "+1 (930) 434-2394", + "address": "786 Polhemus Place, Dale, New Jersey, 1217", + "registered": "2015-07-01T11:34:42 -02:00", + "latitude": 40.193945, + "longitude": 45.243376, + "tags": [ + "aute", + "culpa", + "tempor", + "in", + "ea", + "magna", + "exercitation" + ], + "friends": [ + { + "id": 0, + "name": "Verna Franco" + }, + { + "id": 1, + "name": "Judith Mayer" + }, + { + "id": 2, + "name": "Louisa Blackwell" + } + ], + "greeting": "Hello, Julie Gregory! You have 2 unread messages.", + "favoriteFruit": "apple" + }, + { + "_id": "58fa646d8d5448ee28412568", + "index": 12, + "guid": "5383dc29-7190-4187-8980-a2d4ff78e80e", + "isActive": true, + "balance": "$3,442.36", + "picture": "http://placehold.it/32x32", + "age": 23, + "eyeColor": "green", + "name": "Cathryn Ball", + "gender": "female", + "company": "PUSHCART", + "email": "cathrynball@pushcart.com", + "phone": "+1 (914) 561-3857", + "address": "954 Imlay Street, Gilmore, Nebraska, 9546", + "registered": "2014-09-21T08:44:19 -02:00", + "latitude": -29.965682, + "longitude": -138.147241, + "tags": [ + "occaecat", + "reprehenderit", + "est", + "sint", + "ut", + "consequat", + "duis" + ], + "friends": [ + { + "id": 0, + "name": "Dean Fisher" + }, + { + "id": 1, + "name": "Kris Bowen" + }, + { + "id": 2, + "name": "Orr Baird" + } + ], + "greeting": "Hello, Cathryn Ball! You have 1 unread messages.", + "favoriteFruit": "strawberry" + }, + { + "_id": "58fa646d817d78dc856b5bd8", + "index": 13, + "guid": "61fbb768-6185-4316-a14f-3e2ffed9d7a5", + "isActive": false, + "balance": "$1,409.91", + "picture": "http://placehold.it/32x32", + "age": 21, + "eyeColor": "green", + "name": "Antonia Castaneda", + "gender": "female", + "company": "DANCERITY", + "email": "antoniacastaneda@dancerity.com", + "phone": "+1 (827) 581-2620", + "address": "994 Revere Place, Freetown, Nevada, 7526", + "registered": "2014-09-18T07:40:52 -02:00", + "latitude": 79.298804, + "longitude": 38.173708, + "tags": [ + "pariatur", + "sit", + "elit", + "excepteur", + "enim", + "ut", + "exercitation" + ], + "friends": [ + { + "id": 0, + "name": "Morrow Lane" + }, + { + "id": 1, + "name": "Donovan Malone" + }, + { + "id": 2, + "name": "Anderson Combs" + } + ], + "greeting": "Hello, Antonia Castaneda! You have 10 unread messages.", + "favoriteFruit": "apple" + }, + { + "_id": "58fa646d838c168705f38fb9", + "index": 14, + "guid": "9312a1d4-b19a-43d8-b376-cbafd205af08", + "isActive": true, + "balance": "$3,911.02", + "picture": "http://placehold.it/32x32", + "age": 20, + "eyeColor": "brown", + "name": "Hardy Hooper", + "gender": "male", + "company": "AQUAZURE", + "email": "hardyhooper@aquazure.com", + "phone": "+1 (960) 485-3693", + "address": "598 Hanson Place, Mammoth, Mississippi, 4764", + "registered": "2016-05-09T03:07:37 -02:00", + "latitude": 47.54255, + "longitude": -93.908637, + "tags": [ + "nulla", + "deserunt", + "dolore", + "ipsum", + "cillum", + "exercitation", + "pariatur" + ], + "friends": [ + { + "id": 0, + "name": "Ramona Underwood" + }, + { + "id": 1, + "name": "Trudy Reid" + }, + { + "id": 2, + "name": "Nettie Talley" + } + ], + "greeting": "Hello, Hardy Hooper! You have 5 unread messages.", + "favoriteFruit": "banana" + }, + { + "_id": "58fa646d905949740fe2965e", + "index": 15, + "guid": "d74f0592-7218-4fe1-9744-e83759a51596", + "isActive": true, + "balance": "$3,384.05", + "picture": "http://placehold.it/32x32", + "age": 29, + "eyeColor": "blue", + "name": "Concetta Simmons", + "gender": "female", + "company": "PODUNK", + "email": "concettasimmons@podunk.com", + "phone": "+1 (982) 535-2748", + "address": "496 Waldane Court, Charco, New York, 5806", + "registered": "2015-02-16T02:54:01 -01:00", + "latitude": 47.850051, + "longitude": 165.157751, + "tags": [ + "Lorem", + "exercitation", + "enim", + "aliqua", + "id", + "laboris", + "officia" + ], + "friends": [ + { + "id": 0, + "name": "Tara Warren" + }, + { + "id": 1, + "name": "Leanne Rodgers" + }, + { + "id": 2, + "name": "Stacey Thornton" + } + ], + "greeting": "Hello, Concetta Simmons! You have 1 unread messages.", + "favoriteFruit": "apple" + }, + { + "_id": "58fa646de536a4cbdee0073e", + "index": 16, + "guid": "f15a0f38-721d-4a1c-a960-37df0f2d97cd", + "isActive": true, + "balance": "$2,471.47", + "picture": "http://placehold.it/32x32", + "age": 36, + "eyeColor": "brown", + "name": "Aurora Griffin", + "gender": "female", + "company": "ENTROFLEX", + "email": "auroragriffin@entroflex.com", + "phone": "+1 (969) 527-2632", + "address": "587 Vanderveer Street, Bentonville, Oklahoma, 1715", + "registered": "2017-01-13T10:22:35 -01:00", + "latitude": -18.780868, + "longitude": -73.008945, + "tags": [ + "nisi", + "laborum", + "labore", + "irure", + "commodo", + "ex", + "dolor" + ], + "friends": [ + { + "id": 0, + "name": "Lamb Perez" + }, + { + "id": 1, + "name": "Linda Fletcher" + }, + { + "id": 2, + "name": "Gutierrez Durham" + } + ], + "greeting": "Hello, Aurora Griffin! You have 5 unread messages.", + "favoriteFruit": "banana" + }, + { + "_id": "58fa646d0b5c5dcd846310a6", + "index": 17, + "guid": "d9d91e6e-2b7b-46fd-9db5-68f93e389d04", + "isActive": true, + "balance": "$1,165.95", + "picture": "http://placehold.it/32x32", + "age": 21, + "eyeColor": "brown", + "name": "Tricia Blake", + "gender": "female", + "company": "RAMJOB", + "email": "triciablake@ramjob.com", + "phone": "+1 (995) 440-2261", + "address": "595 Pioneer Street, Greer, Maryland, 3958", + "registered": "2014-03-28T03:52:56 -01:00", + "latitude": -43.601892, + "longitude": -51.434455, + "tags": [ + "mollit", + "pariatur", + "nostrud", + "proident", + "aliqua", + "pariatur", + "sit" + ], + "friends": [ + { + "id": 0, + "name": "Mays Roach" + }, + { + "id": 1, + "name": "Concepcion Whitaker" + }, + { + "id": 2, + "name": "Cardenas Warner" + } + ], + "greeting": "Hello, Tricia Blake! You have 8 unread messages.", + "favoriteFruit": "banana" + }, + { + "_id": "58fa646d6e56ed13e165a7bc", + "index": 18, + "guid": "570984dd-a069-47bb-93a3-65fabd81c11a", + "isActive": true, + "balance": "$2,800.89", + "picture": "http://placehold.it/32x32", + "age": 33, + "eyeColor": "blue", + "name": "Buck Oneal", + "gender": "male", + "company": "EXOSIS", + "email": "buckoneal@exosis.com", + "phone": "+1 (995) 548-3851", + "address": "954 Johnson Avenue, Sanders, Wyoming, 8101", + "registered": "2014-04-01T12:15:47 -02:00", + "latitude": 50.843966, + "longitude": 9.807391, + "tags": [ + "id", + "ipsum", + "consequat", + "et", + "amet", + "cillum", + "est" + ], + "friends": [ + { + "id": 0, + "name": "Tameka Martinez" + }, + { + "id": 1, + "name": "Hammond Frank" + }, + { + "id": 2, + "name": "Nora Hodge" + } + ], + "greeting": "Hello, Buck Oneal! You have 8 unread messages.", + "favoriteFruit": "banana" + }, + { + "_id": "58fa646d65372dfec7c43fe1", + "index": 19, + "guid": "74b86aba-c9d1-4559-a445-13254b62187c", + "isActive": true, + "balance": "$1,276.12", + "picture": "http://placehold.it/32x32", + "age": 36, + "eyeColor": "brown", + "name": "Galloway Clayton", + "gender": "male", + "company": "COASH", + "email": "gallowayclayton@coash.com", + "phone": "+1 (898) 504-2401", + "address": "930 Everit Street, Hiseville, Arizona, 2510", + "registered": "2016-03-02T02:25:08 -01:00", + "latitude": -12.170306, + "longitude": -95.908687, + "tags": [ + "aute", + "esse", + "irure", + "qui", + "excepteur", + "Lorem", + "laboris" + ], + "friends": [ + { + "id": 0, + "name": "Shauna Boone" + }, + { + "id": 1, + "name": "Jamie Hurst" + }, + { + "id": 2, + "name": "Erna Francis" + } + ], + "greeting": "Hello, Galloway Clayton! You have 7 unread messages.", + "favoriteFruit": "banana" + }, + { + "_id": "58fa646d2319a3722c2577e3", + "index": 20, + "guid": "40c7ff40-74ac-4266-8f01-a15eb4f887bc", + "isActive": true, + "balance": "$3,922.67", + "picture": "http://placehold.it/32x32", + "age": 38, + "eyeColor": "blue", + "name": "Flora Guzman", + "gender": "female", + "company": "REMOTION", + "email": "floraguzman@remotion.com", + "phone": "+1 (990) 551-2830", + "address": "516 Portland Avenue, Lawrence, Texas, 8562", + "registered": "2014-12-11T09:11:54 -01:00", + "latitude": 76.309479, + "longitude": -84.795771, + "tags": [ + "anim", + "irure", + "culpa", + "nisi", + "tempor", + "cillum", + "sit" + ], + "friends": [ + { + "id": 0, + "name": "Bettie Wagner" + }, + { + "id": 1, + "name": "Cathleen Blevins" + }, + { + "id": 2, + "name": "Ina Lewis" + } + ], + "greeting": "Hello, Flora Guzman! You have 7 unread messages.", + "favoriteFruit": "strawberry" + }, + { + "_id": "58fa646d73ce4a3db223451c", + "index": 21, + "guid": "f478675a-982d-44b6-9496-bd441d9b9150", + "isActive": true, + "balance": "$3,484.58", + "picture": "http://placehold.it/32x32", + "age": 40, + "eyeColor": "brown", + "name": "Camacho Hull", + "gender": "male", + "company": "EARTHMARK", + "email": "camachohull@earthmark.com", + "phone": "+1 (939) 441-3262", + "address": "892 Euclid Avenue, Cochranville, Connecticut, 566", + "registered": "2015-02-13T11:41:23 -01:00", + "latitude": -7.525654, + "longitude": -137.227222, + "tags": [ + "velit", + "ea", + "dolor", + "consectetur", + "fugiat", + "magna", + "deserunt" + ], + "friends": [ + { + "id": 0, + "name": "Stone Ruiz" + }, + { + "id": 1, + "name": "Estela Reese" + }, + { + "id": 2, + "name": "Carroll Chandler" + } + ], + "greeting": "Hello, Camacho Hull! You have 5 unread messages.", + "favoriteFruit": "apple" + }, + { + "_id": "58fa646d44060c1d29e188a3", + "index": 22, + "guid": "6f875c3f-7d2f-4b2d-906d-f1ab66dd09bc", + "isActive": false, + "balance": "$2,003.48", + "picture": "http://placehold.it/32x32", + "age": 22, + "eyeColor": "blue", + "name": "Marks Lucas", + "gender": "male", + "company": "ZOLARITY", + "email": "markslucas@zolarity.com", + "phone": "+1 (842) 536-2517", + "address": "985 Covert Street, Winston, Maine, 616", + "registered": "2015-01-31T04:28:19 -01:00", + "latitude": 74.390901, + "longitude": -166.116826, + "tags": [ + "adipisicing", + "aliqua", + "eu", + "occaecat", + "ad", + "anim", + "cillum" + ], + "friends": [ + { + "id": 0, + "name": "Sawyer Guerrero" + }, + { + "id": 1, + "name": "Cortez Craft" + }, + { + "id": 2, + "name": "Juliette Ballard" + } + ], + "greeting": "Hello, Marks Lucas! You have 9 unread messages.", + "favoriteFruit": "apple" + }, + { + "_id": "58fa646d94bdee2b8174bbf3", + "index": 23, + "guid": "18904cce-8014-4d9e-8146-7aaa42e008ca", + "isActive": true, + "balance": "$1,636.46", + "picture": "http://placehold.it/32x32", + "age": 22, + "eyeColor": "blue", + "name": "Horne Hurley", + "gender": "male", + "company": "AUSTEX", + "email": "hornehurley@austex.com", + "phone": "+1 (850) 558-3496", + "address": "916 Blake Court, Fidelis, Idaho, 8161", + "registered": "2015-12-23T12:33:02 -01:00", + "latitude": 71.944251, + "longitude": 35.479952, + "tags": [ + "excepteur", + "laboris", + "qui", + "commodo", + "esse", + "tempor", + "dolor" + ], + "friends": [ + { + "id": 0, + "name": "Briana Lowe" + }, + { + "id": 1, + "name": "Byrd William" + }, + { + "id": 2, + "name": "Lilia Johnston" + } + ], + "greeting": "Hello, Horne Hurley! You have 7 unread messages.", + "favoriteFruit": "apple" + }, + { + "_id": "58fa646dd44d8a45556b888a", + "index": 24, + "guid": "e2907f86-aad5-4753-9904-bea1607f4a7d", + "isActive": true, + "balance": "$1,582.27", + "picture": "http://placehold.it/32x32", + "age": 32, + "eyeColor": "blue", + "name": "Kerr Ray", + "gender": "male", + "company": "KIOSK", + "email": "kerrray@kiosk.com", + "phone": "+1 (955) 564-3568", + "address": "363 Poly Place, Hall, Hawaii, 9178", + "registered": "2015-11-27T11:56:35 -01:00", + "latitude": 89.768612, + "longitude": 155.288508, + "tags": [ + "dolor", + "quis", + "nulla", + "proident", + "in", + "id", + "dolor" + ], + "friends": [ + { + "id": 0, + "name": "Irwin Sellers" + }, + { + "id": 1, + "name": "Calhoun Sharpe" + }, + { + "id": 2, + "name": "Earnestine Sherman" + } + ], + "greeting": "Hello, Kerr Ray! You have 4 unread messages.", + "favoriteFruit": "apple" + }, + { + "_id": "58fa646d61a77141b61eb9a7", + "index": 25, + "guid": "d8aadb97-8050-491e-8b32-115a9710b37c", + "isActive": true, + "balance": "$2,017.90", + "picture": "http://placehold.it/32x32", + "age": 34, + "eyeColor": "blue", + "name": "Rosales Shepherd", + "gender": "male", + "company": "ESCHOIR", + "email": "rosalesshepherd@eschoir.com", + "phone": "+1 (919) 598-2946", + "address": "296 Coffey Street, Lookingglass, Alaska, 3898", + "registered": "2015-12-11T04:47:58 -01:00", + "latitude": -53.495093, + "longitude": -140.310969, + "tags": [ + "ad", + "elit", + "ea", + "consectetur", + "veniam", + "est", + "tempor" + ], + "friends": [ + { + "id": 0, + "name": "Savannah Jones" + }, + { + "id": 1, + "name": "Gale Oliver" + }, + { + "id": 2, + "name": "Rice Mullins" + } + ], + "greeting": "Hello, Rosales Shepherd! You have 9 unread messages.", + "favoriteFruit": "apple" + }, + { + "_id": "58fa646df41443122560b954", + "index": 26, + "guid": "4fc501cf-3af2-4c44-816d-5257f8e59109", + "isActive": false, + "balance": "$1,639.48", + "picture": "http://placehold.it/32x32", + "age": 21, + "eyeColor": "green", + "name": "Chris Cline", + "gender": "female", + "company": "AQUACINE", + "email": "chriscline@aquacine.com", + "phone": "+1 (931) 524-3562", + "address": "470 Cheever Place, Lopezo, Iowa, 6027", + "registered": "2014-04-23T09:34:53 -02:00", + "latitude": -77.009748, + "longitude": -56.275735, + "tags": [ + "minim", + "deserunt", + "magna", + "elit", + "non", + "adipisicing", + "duis" + ], + "friends": [ + { + "id": 0, + "name": "Marion Campos" + }, + { + "id": 1, + "name": "Geneva Donaldson" + }, + { + "id": 2, + "name": "Madden Ellis" + } + ], + "greeting": "Hello, Chris Cline! You have 2 unread messages.", + "favoriteFruit": "banana" + }, + { + "_id": "58fa646d6f336367eadcae58", + "index": 27, + "guid": "a46401ee-46d0-4dba-b7df-fb63429edb4b", + "isActive": true, + "balance": "$3,501.55", + "picture": "http://placehold.it/32x32", + "age": 36, + "eyeColor": "blue", + "name": "Parker French", + "gender": "male", + "company": "GEEKUS", + "email": "parkerfrench@geekus.com", + "phone": "+1 (929) 418-3732", + "address": "902 Abbey Court, Convent, Michigan, 9373", + "registered": "2017-01-04T07:18:21 -01:00", + "latitude": 89.576429, + "longitude": -62.68681, + "tags": [ + "et", + "incididunt", + "excepteur", + "proident", + "in", + "do", + "elit" + ], + "friends": [ + { + "id": 0, + "name": "Yolanda Silva" + }, + { + "id": 1, + "name": "Mason Meyer" + }, + { + "id": 2, + "name": "Yates Orr" + } + ], + "greeting": "Hello, Parker French! You have 3 unread messages.", + "favoriteFruit": "apple" + }, + { + "_id": "58fa646d8b89360f427944c0", + "index": 28, + "guid": "59d532b7-3560-4185-ac21-c48efd88ca37", + "isActive": false, + "balance": "$2,158.03", + "picture": "http://placehold.it/32x32", + "age": 35, + "eyeColor": "brown", + "name": "Saunders Stevenson", + "gender": "male", + "company": "ACRUEX", + "email": "saundersstevenson@acruex.com", + "phone": "+1 (828) 574-2096", + "address": "982 Mill Avenue, Turpin, North Dakota, 5081", + "registered": "2017-03-20T03:08:17 -01:00", + "latitude": 80.821621, + "longitude": -122.267328, + "tags": [ + "voluptate", + "qui", + "et", + "esse", + "sint", + "velit", + "esse" + ], + "friends": [ + { + "id": 0, + "name": "Margarita Waller" + }, + { + "id": 1, + "name": "Julia Barton" + }, + { + "id": 2, + "name": "Kitty Trujillo" + } + ], + "greeting": "Hello, Saunders Stevenson! You have 1 unread messages.", + "favoriteFruit": "banana" + }, + { + "_id": "58fa646d904a8f1113deac3d", + "index": 29, + "guid": "48a5fe7e-fa66-494a-8e96-432bbe70bda0", + "isActive": true, + "balance": "$3,621.35", + "picture": "http://placehold.it/32x32", + "age": 40, + "eyeColor": "green", + "name": "Burnett Phillips", + "gender": "male", + "company": "MYOPIUM", + "email": "burnettphillips@myopium.com", + "phone": "+1 (902) 529-2879", + "address": "179 College Place, Kingstowne, Guam, 9184", + "registered": "2016-02-29T07:46:42 -01:00", + "latitude": -59.153019, + "longitude": -49.35699, + "tags": [ + "magna", + "laboris", + "ut", + "nisi", + "sunt", + "sint", + "sunt" + ], + "friends": [ + { + "id": 0, + "name": "Mcdowell Hebert" + }, + { + "id": 1, + "name": "Rodriguez Gonzales" + }, + { + "id": 2, + "name": "Casandra Ferrell" + } + ], + "greeting": "Hello, Burnett Phillips! You have 9 unread messages.", + "favoriteFruit": "banana" + }, + { + "_id": "58fa646d9347947c735a74bd", + "index": 30, + "guid": "59b17a98-b5fb-4e69-9cfd-93f03a862f35", + "isActive": false, + "balance": "$2,129.93", + "picture": "http://placehold.it/32x32", + "age": 34, + "eyeColor": "brown", + "name": "Wiley Riddle", + "gender": "male", + "company": "ASSISTIX", + "email": "wileyriddle@assistix.com", + "phone": "+1 (930) 568-2289", + "address": "914 Bayard Street, Beaverdale, Virginia, 542", + "registered": "2015-02-02T02:22:50 -01:00", + "latitude": -59.2455, + "longitude": -124.569698, + "tags": [ + "dolor", + "est", + "irure", + "eiusmod", + "ipsum", + "aliqua", + "enim" + ], + "friends": [ + { + "id": 0, + "name": "Pate Holder" + }, + { + "id": 1, + "name": "Eugenia Valentine" + }, + { + "id": 2, + "name": "Luisa Cobb" + } + ], + "greeting": "Hello, Wiley Riddle! You have 4 unread messages.", + "favoriteFruit": "banana" + }, + { + "_id": "58fa646da47b620a64f8dd27", + "index": 31, + "guid": "1a767d7b-a93b-4d7c-8b55-45f2fb9d1cc9", + "isActive": true, + "balance": "$1,822.86", + "picture": "http://placehold.it/32x32", + "age": 35, + "eyeColor": "brown", + "name": "Rosie Greer", + "gender": "female", + "company": "XERONK", + "email": "rosiegreer@xeronk.com", + "phone": "+1 (959) 443-2381", + "address": "336 Lawton Street, Finderne, West Virginia, 4904", + "registered": "2016-05-26T04:31:34 -02:00", + "latitude": 1.482268, + "longitude": 5.585765, + "tags": [ + "quis", + "exercitation", + "consequat", + "sunt", + "sit", + "est", + "sunt" + ], + "friends": [ + { + "id": 0, + "name": "Aline Stone" + }, + { + "id": 1, + "name": "Chambers Pope" + }, + { + "id": 2, + "name": "Morales Langley" + } + ], + "greeting": "Hello, Rosie Greer! You have 2 unread messages.", + "favoriteFruit": "strawberry" + }, + { + "_id": "58fa646d950b7a28dc96960d", + "index": 32, + "guid": "605c9733-eeb4-474b-a548-678daa5ff51b", + "isActive": true, + "balance": "$3,619.58", + "picture": "http://placehold.it/32x32", + "age": 35, + "eyeColor": "blue", + "name": "Velez Sheppard", + "gender": "male", + "company": "ENQUILITY", + "email": "velezsheppard@enquility.com", + "phone": "+1 (897) 502-3600", + "address": "633 Montgomery Street, Wintersburg, Kentucky, 7260", + "registered": "2015-01-09T09:53:13 -01:00", + "latitude": 11.071945, + "longitude": 16.982668, + "tags": [ + "officia", + "velit", + "aute", + "nulla", + "duis", + "incididunt", + "proident" + ], + "friends": [ + { + "id": 0, + "name": "Reid Melendez" + }, + { + "id": 1, + "name": "Mcdaniel Perry" + }, + { + "id": 2, + "name": "Peggy Ellison" + } + ], + "greeting": "Hello, Velez Sheppard! You have 10 unread messages.", + "favoriteFruit": "strawberry" + }, + { + "_id": "58fa646d66d1a50fbb02dd6e", + "index": 33, + "guid": "f3ba0a23-7643-49a9-83a4-c24af02ea0d8", + "isActive": true, + "balance": "$2,848.94", + "picture": "http://placehold.it/32x32", + "age": 20, + "eyeColor": "green", + "name": "Iris Mills", + "gender": "female", + "company": "IMKAN", + "email": "irismills@imkan.com", + "phone": "+1 (994) 441-3633", + "address": "562 Seba Avenue, Ticonderoga, Northern Mariana Islands, 2485", + "registered": "2014-02-21T03:29:03 -01:00", + "latitude": -38.968902, + "longitude": 80.200156, + "tags": [ + "cillum", + "minim", + "fugiat", + "enim", + "irure", + "magna", + "eu" + ], + "friends": [ + { + "id": 0, + "name": "Hayes James" + }, + { + "id": 1, + "name": "Fuller Pollard" + }, + { + "id": 2, + "name": "Thomas Cox" + } + ], + "greeting": "Hello, Iris Mills! You have 5 unread messages.", + "favoriteFruit": "strawberry" + }, + { + "_id": "58fa646d6e224718c0c54cf9", + "index": 34, + "guid": "da735ea6-47b9-48bf-b2c3-e73323d7f66c", + "isActive": false, + "balance": "$1,188.71", + "picture": "http://placehold.it/32x32", + "age": 40, + "eyeColor": "brown", + "name": "Judy Harper", + "gender": "female", + "company": "ANIXANG", + "email": "judyharper@anixang.com", + "phone": "+1 (830) 562-3810", + "address": "553 Newkirk Placez, Condon, Virgin Islands, 8147", + "registered": "2014-07-23T01:45:32 -02:00", + "latitude": -20.134894, + "longitude": 12.073907, + "tags": [ + "cillum", + "tempor", + "officia", + "deserunt", + "do", + "duis", + "reprehenderit" + ], + "friends": [ + { + "id": 0, + "name": "Joy Richard" + }, + { + "id": 1, + "name": "Jenna Duncan" + }, + { + "id": 2, + "name": "Krystal Vinson" + } + ], + "greeting": "Hello, Judy Harper! You have 7 unread messages.", + "favoriteFruit": "banana" + }, + { + "_id": "58fa646de718be5ec45055c5", + "index": 35, + "guid": "39a73a7b-6c80-4aff-97ed-17c431b028d5", + "isActive": false, + "balance": "$3,249.21", + "picture": "http://placehold.it/32x32", + "age": 38, + "eyeColor": "green", + "name": "Wise Bell", + "gender": "male", + "company": "EXOVENT", + "email": "wisebell@exovent.com", + "phone": "+1 (961) 435-2204", + "address": "346 Everett Avenue, Foxworth, Federated States Of Micronesia, 7880", + "registered": "2014-01-29T12:50:10 -01:00", + "latitude": -83.382857, + "longitude": -85.641786, + "tags": [ + "Lorem", + "cillum", + "Lorem", + "nulla", + "consequat", + "mollit", + "ad" + ], + "friends": [ + { + "id": 0, + "name": "Wilma Hyde" + }, + { + "id": 1, + "name": "Lynn Stokes" + }, + { + "id": 2, + "name": "Mcgowan Oneill" + } + ], + "greeting": "Hello, Wise Bell! You have 7 unread messages.", + "favoriteFruit": "banana" + }, + { + "_id": "58fa646d2fe7f5b98359e772", + "index": 36, + "guid": "3e4ba71d-de00-4573-b887-a505b3721b70", + "isActive": false, + "balance": "$1,978.68", + "picture": "http://placehold.it/32x32", + "age": 36, + "eyeColor": "blue", + "name": "Annette Davis", + "gender": "female", + "company": "KONGENE", + "email": "annettedavis@kongene.com", + "phone": "+1 (837) 546-3401", + "address": "268 Centre Street, Otranto, Utah, 3698", + "registered": "2016-09-02T10:26:49 -02:00", + "latitude": -43.029939, + "longitude": -1.397371, + "tags": [ + "ad", + "laboris", + "non", + "proident", + "tempor", + "exercitation", + "nisi" + ], + "friends": [ + { + "id": 0, + "name": "Melisa Rojas" + }, + { + "id": 1, + "name": "Langley Dean" + }, + { + "id": 2, + "name": "Ebony Mack" + } + ], + "greeting": "Hello, Annette Davis! You have 5 unread messages.", + "favoriteFruit": "apple" + }, + { + "_id": "58fa646d9ad29d932e02d521", + "index": 37, + "guid": "a82404fe-4512-437b-a711-9f4d27c6d9aa", + "isActive": false, + "balance": "$2,099.63", + "picture": "http://placehold.it/32x32", + "age": 24, + "eyeColor": "brown", + "name": "Felicia Woodward", + "gender": "female", + "company": "MITROC", + "email": "feliciawoodward@mitroc.com", + "phone": "+1 (967) 424-2479", + "address": "362 Navy Walk, Celeryville, Georgia, 3547", + "registered": "2014-07-04T03:19:57 -02:00", + "latitude": 69.903583, + "longitude": -178.989099, + "tags": [ + "laboris", + "est", + "est", + "nulla", + "commodo", + "non", + "laboris" + ], + "friends": [ + { + "id": 0, + "name": "Sally Monroe" + }, + { + "id": 1, + "name": "Kerry Howell" + }, + { + "id": 2, + "name": "Kinney Brewer" + } + ], + "greeting": "Hello, Felicia Woodward! You have 10 unread messages.", + "favoriteFruit": "banana" + }, + { + "_id": "58fa646d68894fa2885955b6", + "index": 38, + "guid": "3380be4a-382d-4722-bf25-519f9730fc96", + "isActive": true, + "balance": "$3,699.36", + "picture": "http://placehold.it/32x32", + "age": 28, + "eyeColor": "green", + "name": "Manning Waters", + "gender": "male", + "company": "ENVIRE", + "email": "manningwaters@envire.com", + "phone": "+1 (921) 421-2487", + "address": "169 Bradford Street, Juarez, Pennsylvania, 2988", + "registered": "2015-04-12T12:46:26 -02:00", + "latitude": 13.246711, + "longitude": -74.68705, + "tags": [ + "consequat", + "magna", + "consequat", + "culpa", + "proident", + "velit", + "aliquip" + ], + "friends": [ + { + "id": 0, + "name": "Christi West" + }, + { + "id": 1, + "name": "Imogene Wooten" + }, + { + "id": 2, + "name": "Carey Vega" + } + ], + "greeting": "Hello, Manning Waters! You have 10 unread messages.", + "favoriteFruit": "apple" + }, + { + "_id": "58fa646d5134a659f923d54c", + "index": 39, + "guid": "947b42ec-e12c-4d9e-a285-1104f86607af", + "isActive": false, + "balance": "$3,772.51", + "picture": "http://placehold.it/32x32", + "age": 40, + "eyeColor": "blue", + "name": "Sheryl Walls", + "gender": "female", + "company": "COMTRAIL", + "email": "sherylwalls@comtrail.com", + "phone": "+1 (998) 444-3442", + "address": "613 Fair Street, Worton, Indiana, 7308", + "registered": "2015-12-16T07:23:13 -01:00", + "latitude": 53.048471, + "longitude": -19.420687, + "tags": [ + "eiusmod", + "duis", + "tempor", + "ut", + "labore", + "amet", + "dolor" + ], + "friends": [ + { + "id": 0, + "name": "Pace Sanford" + }, + { + "id": 1, + "name": "Guadalupe Dalton" + }, + { + "id": 2, + "name": "Nellie Benton" + } + ], + "greeting": "Hello, Sheryl Walls! You have 7 unread messages.", + "favoriteFruit": "banana" + }, + { + "_id": "58fa646d08710b719701a646", + "index": 40, + "guid": "decc524b-d83b-462d-8eb5-3ad694b89de9", + "isActive": true, + "balance": "$2,392.96", + "picture": "http://placehold.it/32x32", + "age": 29, + "eyeColor": "brown", + "name": "Autumn Kelly", + "gender": "female", + "company": "ISOTRACK", + "email": "autumnkelly@isotrack.com", + "phone": "+1 (801) 412-2521", + "address": "730 Nassau Street, Marshall, District Of Columbia, 5453", + "registered": "2014-10-10T04:22:54 -02:00", + "latitude": 11.53084, + "longitude": -151.030545, + "tags": [ + "dolore", + "magna", + "consectetur", + "nostrud", + "fugiat", + "sunt", + "aliqua" + ], + "friends": [ + { + "id": 0, + "name": "Mann Lawrence" + }, + { + "id": 1, + "name": "Torres Sharp" + }, + { + "id": 2, + "name": "Johnnie Guerra" + } + ], + "greeting": "Hello, Autumn Kelly! You have 6 unread messages.", + "favoriteFruit": "apple" + }, + { + "_id": "58fa646d14f672395600ebb4", + "index": 41, + "guid": "abcb6fb6-3b44-4da6-b0dc-8e531c6cddc1", + "isActive": true, + "balance": "$2,795.34", + "picture": "http://placehold.it/32x32", + "age": 33, + "eyeColor": "blue", + "name": "Queen Mckee", + "gender": "female", + "company": "PLASMOS", + "email": "queenmckee@plasmos.com", + "phone": "+1 (862) 473-2493", + "address": "335 Bridge Street, Sunnyside, South Dakota, 6798", + "registered": "2015-03-23T04:39:16 -01:00", + "latitude": -42.142383, + "longitude": 119.914739, + "tags": [ + "culpa", + "irure", + "reprehenderit", + "minim", + "ullamco", + "voluptate", + "ipsum" + ], + "friends": [ + { + "id": 0, + "name": "Dionne Merritt" + }, + { + "id": 1, + "name": "Ivy Gilmore" + }, + { + "id": 2, + "name": "Freda Arnold" + } + ], + "greeting": "Hello, Queen Mckee! You have 10 unread messages.", + "favoriteFruit": "strawberry" + }, + { + "_id": "58fa646dcbecdefae9976b68", + "index": 42, + "guid": "472e6b5f-ae55-4246-8407-35ec6a99d5be", + "isActive": false, + "balance": "$1,333.93", + "picture": "http://placehold.it/32x32", + "age": 40, + "eyeColor": "blue", + "name": "Natalia Landry", + "gender": "female", + "company": "MAGNINA", + "email": "natalialandry@magnina.com", + "phone": "+1 (926) 515-2771", + "address": "562 Hooper Street, Hayes, Ohio, 4843", + "registered": "2016-08-26T10:28:08 -02:00", + "latitude": -46.186231, + "longitude": 139.001522, + "tags": [ + "commodo", + "minim", + "cupidatat", + "non", + "consectetur", + "aute", + "anim" + ], + "friends": [ + { + "id": 0, + "name": "Carlene Chaney" + }, + { + "id": 1, + "name": "Marcie Alford" + }, + { + "id": 2, + "name": "Hines Wilson" + } + ], + "greeting": "Hello, Natalia Landry! You have 10 unread messages.", + "favoriteFruit": "apple" + }, + { + "_id": "58fa646d5bec431f3bcca6ec", + "index": 43, + "guid": "13496d27-e2cc-46ba-806f-c252ca30262c", + "isActive": true, + "balance": "$1,056.14", + "picture": "http://placehold.it/32x32", + "age": 20, + "eyeColor": "blue", + "name": "Amanda Joyce", + "gender": "female", + "company": "OVERFORK", + "email": "amandajoyce@overfork.com", + "phone": "+1 (942) 430-3158", + "address": "433 Aviation Road, Caledonia, Palau, 3348", + "registered": "2017-01-05T05:49:10 -01:00", + "latitude": -35.965417, + "longitude": -149.921618, + "tags": [ + "nulla", + "excepteur", + "exercitation", + "sit", + "labore", + "culpa", + "adipisicing" + ], + "friends": [ + { + "id": 0, + "name": "Carly Mathews" + }, + { + "id": 1, + "name": "Sheila Bishop" + }, + { + "id": 2, + "name": "Deleon Christian" + } + ], + "greeting": "Hello, Amanda Joyce! You have 9 unread messages.", + "favoriteFruit": "strawberry" + }, + { + "_id": "58fa646d72de9b695e707259", + "index": 44, + "guid": "e2401c45-7718-43c4-9251-490b5cc66e69", + "isActive": true, + "balance": "$2,298.68", + "picture": "http://placehold.it/32x32", + "age": 38, + "eyeColor": "brown", + "name": "Stark Conrad", + "gender": "male", + "company": "COMVERGES", + "email": "starkconrad@comverges.com", + "phone": "+1 (843) 535-2224", + "address": "579 Kensington Street, Soham, Louisiana, 9423", + "registered": "2015-12-13T06:34:07 -01:00", + "latitude": -26.563554, + "longitude": -111.834799, + "tags": [ + "proident", + "veniam", + "cillum", + "labore", + "do", + "Lorem", + "quis" + ], + "friends": [ + { + "id": 0, + "name": "Tonya Lott" + }, + { + "id": 1, + "name": "Tanya Fields" + }, + { + "id": 2, + "name": "Kay Espinoza" + } + ], + "greeting": "Hello, Stark Conrad! You have 7 unread messages.", + "favoriteFruit": "strawberry" + }, + { + "_id": "58fa646d21d9e466a949ded4", + "index": 45, + "guid": "1c730823-7d4c-4138-8d63-12699c4e3130", + "isActive": true, + "balance": "$3,728.31", + "picture": "http://placehold.it/32x32", + "age": 24, + "eyeColor": "green", + "name": "Eaton Dodson", + "gender": "male", + "company": "LYRICHORD", + "email": "eatondodson@lyrichord.com", + "phone": "+1 (980) 403-2258", + "address": "472 Woodruff Avenue, Trucksville, Montana, 4090", + "registered": "2016-05-11T12:53:24 -02:00", + "latitude": 52.792622, + "longitude": -36.676171, + "tags": [ + "aliquip", + "do", + "duis", + "sit", + "cupidatat", + "proident", + "proident" + ], + "friends": [ + { + "id": 0, + "name": "Dixie Brady" + }, + { + "id": 1, + "name": "Hurley Bridges" + }, + { + "id": 2, + "name": "Muriel Hudson" + } + ], + "greeting": "Hello, Eaton Dodson! You have 9 unread messages.", + "favoriteFruit": "strawberry" + }, + { + "_id": "58fa646d566853272c6f7875", + "index": 46, + "guid": "751fb288-95cb-47c4-b450-36e8d8b87519", + "isActive": true, + "balance": "$3,006.68", + "picture": "http://placehold.it/32x32", + "age": 25, + "eyeColor": "green", + "name": "Rosario Farley", + "gender": "male", + "company": "KEENGEN", + "email": "rosariofarley@keengen.com", + "phone": "+1 (965) 472-3606", + "address": "459 Richmond Street, Adamstown, Illinois, 8028", + "registered": "2015-06-11T08:04:48 -02:00", + "latitude": 79.124466, + "longitude": 27.645295, + "tags": [ + "nulla", + "eu", + "aliquip", + "tempor", + "et", + "sint", + "sint" + ], + "friends": [ + { + "id": 0, + "name": "Lynne Ewing" + }, + { + "id": 1, + "name": "Joyner Dillon" + }, + { + "id": 2, + "name": "Bass Robles" + } + ], + "greeting": "Hello, Rosario Farley! You have 3 unread messages.", + "favoriteFruit": "strawberry" + }, + { + "_id": "58fa646d8076c092d5429bf5", + "index": 47, + "guid": "14f85652-2d73-43a8-af64-dd7fb6997db5", + "isActive": false, + "balance": "$1,857.44", + "picture": "http://placehold.it/32x32", + "age": 32, + "eyeColor": "blue", + "name": "Leanna Perkins", + "gender": "female", + "company": "XIXAN", + "email": "leannaperkins@xixan.com", + "phone": "+1 (842) 406-3070", + "address": "850 Garden Street, Emory, Minnesota, 5244", + "registered": "2014-12-04T04:12:14 -01:00", + "latitude": -8.176646, + "longitude": 134.038237, + "tags": [ + "occaecat", + "culpa", + "officia", + "officia", + "laboris", + "nulla", + "eu" + ], + "friends": [ + { + "id": 0, + "name": "Shannon Le" + }, + { + "id": 1, + "name": "Mccullough Gardner" + }, + { + "id": 2, + "name": "Mattie Faulkner" + } + ], + "greeting": "Hello, Leanna Perkins! You have 1 unread messages.", + "favoriteFruit": "strawberry" + }, + { + "_id": "58fa646d9f6ef9ac3bad29e6", + "index": 48, + "guid": "95278e02-0d97-44bf-9276-fb135a3d56d8", + "isActive": false, + "balance": "$3,394.10", + "picture": "http://placehold.it/32x32", + "age": 24, + "eyeColor": "brown", + "name": "Estes Reyes", + "gender": "male", + "company": "PARAGONIA", + "email": "estesreyes@paragonia.com", + "phone": "+1 (816) 436-3722", + "address": "831 Granite Street, Hayden, South Carolina, 5725", + "registered": "2014-11-23T04:09:18 -01:00", + "latitude": 36.332002, + "longitude": 70.066886, + "tags": [ + "ea", + "occaecat", + "id", + "exercitation", + "pariatur", + "incididunt", + "pariatur" + ], + "friends": [ + { + "id": 0, + "name": "Gibson Kirby" + }, + { + "id": 1, + "name": "Hickman Hoffman" + }, + { + "id": 2, + "name": "Holden Mendez" + } + ], + "greeting": "Hello, Estes Reyes! You have 9 unread messages.", + "favoriteFruit": "apple" + }, + { + "_id": "58fa646d4cea317886d52b68", + "index": 49, + "guid": "c2896c4a-80b5-4a23-bc89-7bd2851b501d", + "isActive": false, + "balance": "$3,511.65", + "picture": "http://placehold.it/32x32", + "age": 22, + "eyeColor": "brown", + "name": "Rose Wiley", + "gender": "female", + "company": "ACRODANCE", + "email": "rosewiley@acrodance.com", + "phone": "+1 (932) 437-2443", + "address": "915 Jodie Court, Onton, New Mexico, 1085", + "registered": "2014-06-11T09:33:32 -02:00", + "latitude": -73.925502, + "longitude": -25.345141, + "tags": [ + "est", + "quis", + "veniam", + "cillum", + "incididunt", + "tempor", + "fugiat" + ], + "friends": [ + { + "id": 0, + "name": "Irene Andrews" + }, + { + "id": 1, + "name": "Caitlin Rose" + }, + { + "id": 2, + "name": "Erica Payne" + } + ], + "greeting": "Hello, Rose Wiley! You have 1 unread messages.", + "favoriteFruit": "banana" + }, + { + "_id": "58fa646d6a50f303d7dd6485", + "index": 50, + "guid": "3f791160-0c7e-40dd-ab31-5461b261f319", + "isActive": true, + "balance": "$2,245.83", + "picture": "http://placehold.it/32x32", + "age": 25, + "eyeColor": "green", + "name": "Burns Garrison", + "gender": "male", + "company": "INSOURCE", + "email": "burnsgarrison@insource.com", + "phone": "+1 (918) 445-2553", + "address": "974 Lloyd Court, Derwood, North Carolina, 292", + "registered": "2016-12-23T07:41:25 -01:00", + "latitude": 2.181867, + "longitude": 139.026116, + "tags": [ + "laboris", + "excepteur", + "quis", + "aute", + "amet", + "ullamco", + "anim" + ], + "friends": [ + { + "id": 0, + "name": "Davidson Stein" + }, + { + "id": 1, + "name": "Summers Grant" + }, + { + "id": 2, + "name": "Lina Vance" + } + ], + "greeting": "Hello, Burns Garrison! You have 10 unread messages.", + "favoriteFruit": "banana" + }, + { + "_id": "58fa646d1fa30062ffd31bfb", + "index": 51, + "guid": "849255cf-d04e-48f3-9561-2122e7120750", + "isActive": true, + "balance": "$3,431.71", + "picture": "http://placehold.it/32x32", + "age": 40, + "eyeColor": "blue", + "name": "Alissa Williamson", + "gender": "female", + "company": "INTERODEO", + "email": "alissawilliamson@interodeo.com", + "phone": "+1 (950) 409-2688", + "address": "760 Lamont Court, Waterview, Massachusetts, 1112", + "registered": "2015-04-02T07:40:14 -02:00", + "latitude": -57.750754, + "longitude": 56.131544, + "tags": [ + "exercitation", + "ipsum", + "non", + "voluptate", + "labore", + "cillum", + "irure" + ], + "friends": [ + { + "id": 0, + "name": "Sharron Spears" + }, + { + "id": 1, + "name": "Sharpe Mckenzie" + }, + { + "id": 2, + "name": "Lillie Fleming" + } + ], + "greeting": "Hello, Alissa Williamson! You have 3 unread messages.", + "favoriteFruit": "apple" + }, + { + "_id": "58fa646da56b296c890d1b26", + "index": 52, + "guid": "4d89abd3-97de-49ff-a20a-d3dcfca8c71b", + "isActive": false, + "balance": "$3,495.72", + "picture": "http://placehold.it/32x32", + "age": 20, + "eyeColor": "green", + "name": "Patsy Mcdaniel", + "gender": "female", + "company": "BYTREX", + "email": "patsymcdaniel@bytrex.com", + "phone": "+1 (998) 422-2615", + "address": "746 Monroe Place, Bridgetown, Kansas, 6216", + "registered": "2014-09-16T10:43:46 -02:00", + "latitude": 2.099304, + "longitude": -100.374868, + "tags": [ + "esse", + "excepteur", + "mollit", + "veniam", + "adipisicing", + "id", + "sint" + ], + "friends": [ + { + "id": 0, + "name": "Wanda Hendricks" + }, + { + "id": 1, + "name": "Sanchez Carter" + }, + { + "id": 2, + "name": "Imelda Bass" + } + ], + "greeting": "Hello, Patsy Mcdaniel! You have 4 unread messages.", + "favoriteFruit": "strawberry" + }, + { + "_id": "58fa646d76d1d4da8ed51297", + "index": 53, + "guid": "1f4d2532-ee3b-4ef1-93ce-71843df8775d", + "isActive": false, + "balance": "$2,696.38", + "picture": "http://placehold.it/32x32", + "age": 35, + "eyeColor": "green", + "name": "Angie Stephens", + "gender": "female", + "company": "PLUTORQUE", + "email": "angiestephens@plutorque.com", + "phone": "+1 (909) 537-2112", + "address": "772 Dewey Place, Saranap, American Samoa, 4067", + "registered": "2014-03-01T02:16:24 -01:00", + "latitude": 79.906263, + "longitude": -55.812976, + "tags": [ + "fugiat", + "commodo", + "proident", + "magna", + "voluptate", + "reprehenderit", + "sit" + ], + "friends": [ + { + "id": 0, + "name": "Sherry Summers" + }, + { + "id": 1, + "name": "Wilson Stafford" + }, + { + "id": 2, + "name": "Dennis Harris" + } + ], + "greeting": "Hello, Angie Stephens! You have 6 unread messages.", + "favoriteFruit": "banana" + }, + { + "_id": "58fa646d0e034abd3ac2914f", + "index": 54, + "guid": "dda94499-be38-4e24-aa8d-7d0e21d38deb", + "isActive": true, + "balance": "$3,334.18", + "picture": "http://placehold.it/32x32", + "age": 24, + "eyeColor": "blue", + "name": "Margery Page", + "gender": "female", + "company": "TSUNAMIA", + "email": "margerypage@tsunamia.com", + "phone": "+1 (940) 465-2247", + "address": "103 Ryder Avenue, Edmund, Tennessee, 3060", + "registered": "2015-10-17T12:01:42 -02:00", + "latitude": 28.346648, + "longitude": 178.002886, + "tags": [ + "cupidatat", + "deserunt", + "enim", + "aliquip", + "eiusmod", + "dolor", + "mollit" + ], + "friends": [ + { + "id": 0, + "name": "Taylor Powers" + }, + { + "id": 1, + "name": "Christina Blackburn" + }, + { + "id": 2, + "name": "Holt Willis" + } + ], + "greeting": "Hello, Margery Page! You have 4 unread messages.", + "favoriteFruit": "apple" + }, + { + "_id": "58fa646d95516130679cb2f3", + "index": 55, + "guid": "583a6fcf-56aa-453a-aa73-2d66ca007b70", + "isActive": false, + "balance": "$2,684.01", + "picture": "http://placehold.it/32x32", + "age": 36, + "eyeColor": "blue", + "name": "Lopez Cardenas", + "gender": "male", + "company": "MEDMEX", + "email": "lopezcardenas@medmex.com", + "phone": "+1 (896) 550-3416", + "address": "306 Turner Place, Takilma, Florida, 2609", + "registered": "2014-11-10T07:31:35 -01:00", + "latitude": 75.787609, + "longitude": 57.757358, + "tags": [ + "minim", + "pariatur", + "ad", + "in", + "nisi", + "irure", + "pariatur" + ], + "friends": [ + { + "id": 0, + "name": "Latasha Hammond" + }, + { + "id": 1, + "name": "Janine Crane" + }, + { + "id": 2, + "name": "Pearl Beach" + } + ], + "greeting": "Hello, Lopez Cardenas! You have 8 unread messages.", + "favoriteFruit": "apple" + }, + { + "_id": "58fa646d9a7302ffc449e316", + "index": 56, + "guid": "7e1d1a34-4494-4a47-bdfb-2401cf8e6f68", + "isActive": false, + "balance": "$2,624.42", + "picture": "http://placehold.it/32x32", + "age": 38, + "eyeColor": "brown", + "name": "Page Carver", + "gender": "male", + "company": "COMVEYER", + "email": "pagecarver@comveyer.com", + "phone": "+1 (928) 593-2802", + "address": "433 Prince Street, Chapin, Missouri, 8669", + "registered": "2016-11-03T09:08:14 -01:00", + "latitude": 82.307898, + "longitude": 138.379295, + "tags": [ + "mollit", + "cillum", + "laborum", + "enim", + "pariatur", + "fugiat", + "adipisicing" + ], + "friends": [ + { + "id": 0, + "name": "Graves Villarreal" + }, + { + "id": 1, + "name": "Nancy Ingram" + }, + { + "id": 2, + "name": "Loraine Rodriquez" + } + ], + "greeting": "Hello, Page Carver! You have 4 unread messages.", + "favoriteFruit": "strawberry" + }, + { + "_id": "58fa646dcdb4a179818a4bd8", + "index": 57, + "guid": "55d96b3e-bbbe-4060-93e6-6e4987caa3df", + "isActive": true, + "balance": "$3,546.62", + "picture": "http://placehold.it/32x32", + "age": 39, + "eyeColor": "brown", + "name": "Young Carey", + "gender": "female", + "company": "ROCKLOGIC", + "email": "youngcarey@rocklogic.com", + "phone": "+1 (966) 525-2007", + "address": "303 Rose Street, Hemlock, Marshall Islands, 8223", + "registered": "2014-07-13T04:22:06 -02:00", + "latitude": 63.059603, + "longitude": 71.95487, + "tags": [ + "magna", + "dolore", + "voluptate", + "ex", + "consectetur", + "proident", + "duis" + ], + "friends": [ + { + "id": 0, + "name": "Walter Luna" + }, + { + "id": 1, + "name": "Pope Davenport" + }, + { + "id": 2, + "name": "Estrada Mcgowan" + } + ], + "greeting": "Hello, Young Carey! You have 4 unread messages.", + "favoriteFruit": "strawberry" + }, + { + "_id": "58fa646d0f7e11c5ce919e62", + "index": 58, + "guid": "4876e1fd-0915-4357-b5d2-9939d37831fb", + "isActive": true, + "balance": "$2,070.52", + "picture": "http://placehold.it/32x32", + "age": 28, + "eyeColor": "brown", + "name": "Mari Bowman", + "gender": "female", + "company": "ZYTRAX", + "email": "maribowman@zytrax.com", + "phone": "+1 (907) 527-2409", + "address": "540 Hopkins Street, Somerset, Washington, 4593", + "registered": "2017-02-04T08:01:20 -01:00", + "latitude": 9.914386, + "longitude": -114.437464, + "tags": [ + "adipisicing", + "aliquip", + "ea", + "incididunt", + "labore", + "pariatur", + "Lorem" + ], + "friends": [ + { + "id": 0, + "name": "Woodward Chen" + }, + { + "id": 1, + "name": "Rhodes Hardy" + }, + { + "id": 2, + "name": "Cervantes Martin" + } + ], + "greeting": "Hello, Mari Bowman! You have 4 unread messages.", + "favoriteFruit": "banana" + }, + { + "_id": "58fa646d669b4d76f177a519", + "index": 59, + "guid": "679f10df-2335-41fc-8efe-a4514acaf0eb", + "isActive": false, + "balance": "$2,941.16", + "picture": "http://placehold.it/32x32", + "age": 25, + "eyeColor": "green", + "name": "Jeanine Gilliam", + "gender": "female", + "company": "PETICULAR", + "email": "jeaninegilliam@peticular.com", + "phone": "+1 (839) 589-2934", + "address": "430 Varet Street, Wedgewood, Colorado, 4041", + "registered": "2017-02-17T06:35:00 -01:00", + "latitude": -42.397824, + "longitude": -34.299043, + "tags": [ + "consequat", + "non", + "minim", + "tempor", + "id", + "proident", + "Lorem" + ], + "friends": [ + { + "id": 0, + "name": "Becker Giles" + }, + { + "id": 1, + "name": "Aguirre Boyer" + }, + { + "id": 2, + "name": "Lara Glass" + } + ], + "greeting": "Hello, Jeanine Gilliam! You have 9 unread messages.", + "favoriteFruit": "apple" + }, + { + "_id": "58fa646d797606e485074261", + "index": 60, + "guid": "8087164f-3249-4b7b-89e0-cc26b356cbc2", + "isActive": false, + "balance": "$1,819.30", + "picture": "http://placehold.it/32x32", + "age": 21, + "eyeColor": "blue", + "name": "Jodie Haley", + "gender": "female", + "company": "FREAKIN", + "email": "jodiehaley@freakin.com", + "phone": "+1 (960) 555-3075", + "address": "895 Robert Street, Martinsville, Puerto Rico, 958", + "registered": "2014-11-01T02:38:28 -01:00", + "latitude": 36.344446, + "longitude": 23.954717, + "tags": [ + "do", + "adipisicing", + "ullamco", + "in", + "adipisicing", + "mollit", + "veniam" + ], + "friends": [ + { + "id": 0, + "name": "Tabitha Pugh" + }, + { + "id": 1, + "name": "Roth Potter" + }, + { + "id": 2, + "name": "Tyson Hayden" + } + ], + "greeting": "Hello, Jodie Haley! You have 1 unread messages.", + "favoriteFruit": "strawberry" + }, + { + "_id": "58fa646d7bcff11a56bfd3b1", + "index": 61, + "guid": "18d75040-cfa1-4955-aeab-78ca57589ac5", + "isActive": false, + "balance": "$1,915.57", + "picture": "http://placehold.it/32x32", + "age": 21, + "eyeColor": "green", + "name": "Black Snow", + "gender": "male", + "company": "NUTRALAB", + "email": "blacksnow@nutralab.com", + "phone": "+1 (905) 598-3841", + "address": "279 Anthony Street, Garnet, Rhode Island, 9188", + "registered": "2017-02-06T06:19:46 -01:00", + "latitude": -44.972776, + "longitude": 130.271908, + "tags": [ + "enim", + "pariatur", + "magna", + "excepteur", + "dolor", + "ex", + "eu" + ], + "friends": [ + { + "id": 0, + "name": "Bennett Hood" + }, + { + "id": 1, + "name": "Johanna Sullivan" + }, + { + "id": 2, + "name": "Bean Bernard" + } + ], + "greeting": "Hello, Black Snow! You have 4 unread messages.", + "favoriteFruit": "banana" + }, + { + "_id": "58fa646d84937a032d3323e1", + "index": 62, + "guid": "3a9df9d0-40d6-46e5-bfec-454136f27485", + "isActive": false, + "balance": "$2,518.97", + "picture": "http://placehold.it/32x32", + "age": 33, + "eyeColor": "green", + "name": "James Scott", + "gender": "female", + "company": "ROCKABYE", + "email": "jamesscott@rockabye.com", + "phone": "+1 (800) 477-3289", + "address": "761 Frank Court, Gadsden, Alabama, 1738", + "registered": "2015-09-08T08:19:20 -02:00", + "latitude": -64.446182, + "longitude": 41.627025, + "tags": [ + "laboris", + "Lorem", + "sint", + "mollit", + "voluptate", + "enim", + "fugiat" + ], + "friends": [ + { + "id": 0, + "name": "Mcgee Kerr" + }, + { + "id": 1, + "name": "Kristina Hobbs" + }, + { + "id": 2, + "name": "Betty Ford" + } + ], + "greeting": "Hello, James Scott! You have 8 unread messages.", + "favoriteFruit": "banana" + }, + { + "_id": "58fa646ddb183c2830125fc4", + "index": 63, + "guid": "0412f0ac-8c2a-424e-a900-1e7ef4bce4d2", + "isActive": false, + "balance": "$3,036.28", + "picture": "http://placehold.it/32x32", + "age": 35, + "eyeColor": "brown", + "name": "Fox Burke", + "gender": "male", + "company": "ORBOID", + "email": "foxburke@orboid.com", + "phone": "+1 (995) 567-2257", + "address": "490 Benson Avenue, Rosedale, Wisconsin, 2978", + "registered": "2015-07-02T12:30:55 -02:00", + "latitude": -34.873184, + "longitude": -98.050947, + "tags": [ + "laborum", + "sunt", + "eu", + "id", + "voluptate", + "cillum", + "consectetur" + ], + "friends": [ + { + "id": 0, + "name": "Janna Riggs" + }, + { + "id": 1, + "name": "Hamilton Bowers" + }, + { + "id": 2, + "name": "Kasey Hart" + } + ], + "greeting": "Hello, Fox Burke! You have 1 unread messages.", + "favoriteFruit": "apple" + }, + { + "_id": "58fa646deedd68b0596b9fba", + "index": 64, + "guid": "9727cfb9-7deb-4a7d-84dc-24547dcf6cd7", + "isActive": true, + "balance": "$2,539.25", + "picture": "http://placehold.it/32x32", + "age": 22, + "eyeColor": "brown", + "name": "Terry Colon", + "gender": "female", + "company": "SONIQUE", + "email": "terrycolon@sonique.com", + "phone": "+1 (864) 531-2746", + "address": "584 Stockton Street, Waiohinu, California, 5786", + "registered": "2014-02-28T05:04:56 -01:00", + "latitude": -47.695605, + "longitude": -140.996559, + "tags": [ + "officia", + "cillum", + "aliqua", + "aute", + "cillum", + "consequat", + "eiusmod" + ], + "friends": [ + { + "id": 0, + "name": "Dejesus Marsh" + }, + { + "id": 1, + "name": "Angelica Santiago" + }, + { + "id": 2, + "name": "Guthrie Contreras" + } + ], + "greeting": "Hello, Terry Colon! You have 3 unread messages.", + "favoriteFruit": "strawberry" + }, + { + "_id": "58fa646df1cbf4afa247dd26", + "index": 65, + "guid": "62b0aa11-ae42-4a04-8b08-adb6e7f0abba", + "isActive": false, + "balance": "$3,345.24", + "picture": "http://placehold.it/32x32", + "age": 22, + "eyeColor": "brown", + "name": "Angelia Parker", + "gender": "female", + "company": "ZOARERE", + "email": "angeliaparker@zoarere.com", + "phone": "+1 (869) 473-2645", + "address": "994 Stuart Street, Farmington, Arkansas, 3964", + "registered": "2017-03-15T06:24:44 -01:00", + "latitude": -39.023517, + "longitude": -149.270909, + "tags": [ + "mollit", + "reprehenderit", + "aliquip", + "voluptate", + "tempor", + "voluptate", + "adipisicing" + ], + "friends": [ + { + "id": 0, + "name": "Kirk Nieves" + }, + { + "id": 1, + "name": "Jody Salas" + }, + { + "id": 2, + "name": "Berger Dotson" + } + ], + "greeting": "Hello, Angelia Parker! You have 10 unread messages.", + "favoriteFruit": "apple" + }, + { + "_id": "58fa646d87de4350d3af5a9e", + "index": 66, + "guid": "b66088c0-1764-434e-b171-05a6e587c104", + "isActive": true, + "balance": "$2,478.45", + "picture": "http://placehold.it/32x32", + "age": 36, + "eyeColor": "brown", + "name": "Charles Rodriguez", + "gender": "male", + "company": "ISOSWITCH", + "email": "charlesrodriguez@isoswitch.com", + "phone": "+1 (982) 539-2615", + "address": "187 Rutherford Place, Grapeview, Vermont, 2472", + "registered": "2016-06-25T12:46:10 -02:00", + "latitude": -83.848849, + "longitude": 7.986192, + "tags": [ + "dolore", + "Lorem", + "mollit", + "do", + "esse", + "veniam", + "dolore" + ], + "friends": [ + { + "id": 0, + "name": "Lupe England" + }, + { + "id": 1, + "name": "Allison Flowers" + }, + { + "id": 2, + "name": "Garrison Cunningham" + } + ], + "greeting": "Hello, Charles Rodriguez! You have 2 unread messages.", + "favoriteFruit": "banana" + }, + { + "_id": "58fa646d61830652ddf3226e", + "index": 67, + "guid": "455636c7-07bd-4281-9953-8cc05f344466", + "isActive": true, + "balance": "$1,965.16", + "picture": "http://placehold.it/32x32", + "age": 31, + "eyeColor": "brown", + "name": "Coffey Joseph", + "gender": "male", + "company": "GAPTEC", + "email": "coffeyjoseph@gaptec.com", + "phone": "+1 (811) 426-2438", + "address": "641 Coleridge Street, Tivoli, New Hampshire, 1537", + "registered": "2014-01-22T01:52:12 -01:00", + "latitude": -60.778165, + "longitude": -15.368885, + "tags": [ + "pariatur", + "esse", + "aute", + "ipsum", + "occaecat", + "nulla", + "irure" + ], + "friends": [ + { + "id": 0, + "name": "Constance Duran" + }, + { + "id": 1, + "name": "Misty Nunez" + }, + { + "id": 2, + "name": "Mcclure Mcmahon" + } + ], + "greeting": "Hello, Coffey Joseph! You have 9 unread messages.", + "favoriteFruit": "strawberry" + }, + { + "_id": "58fa646d76df67f8e72a7011", + "index": 68, + "guid": "277b5184-c46b-4e7a-a412-9f4873705e5b", + "isActive": true, + "balance": "$1,525.91", + "picture": "http://placehold.it/32x32", + "age": 31, + "eyeColor": "green", + "name": "Moss Santos", + "gender": "male", + "company": "FROSNEX", + "email": "mosssantos@frosnex.com", + "phone": "+1 (896) 441-2511", + "address": "397 Jay Street, Hoagland, Delaware, 1519", + "registered": "2016-06-12T04:07:16 -02:00", + "latitude": -72.308999, + "longitude": -171.272797, + "tags": [ + "exercitation", + "occaecat", + "eu", + "sunt", + "culpa", + "commodo", + "quis" + ], + "friends": [ + { + "id": 0, + "name": "Diana Chan" + }, + { + "id": 1, + "name": "Selena Crosby" + }, + { + "id": 2, + "name": "Mcknight Cummings" + } + ], + "greeting": "Hello, Moss Santos! You have 9 unread messages.", + "favoriteFruit": "banana" + }, + { + "_id": "58fa646db996cc383f9c8999", + "index": 69, + "guid": "09a57810-a6d9-4294-9ea7-cec84eb9b823", + "isActive": true, + "balance": "$3,279.75", + "picture": "http://placehold.it/32x32", + "age": 32, + "eyeColor": "green", + "name": "Keller Valenzuela", + "gender": "male", + "company": "BICOL", + "email": "kellervalenzuela@bicol.com", + "phone": "+1 (832) 587-3173", + "address": "884 Seigel Street, Coral, New Jersey, 666", + "registered": "2016-01-20T07:28:57 -01:00", + "latitude": -37.552671, + "longitude": -152.699635, + "tags": [ + "fugiat", + "consequat", + "eiusmod", + "officia", + "consequat", + "elit", + "ipsum" + ], + "friends": [ + { + "id": 0, + "name": "Tamera Sanders" + }, + { + "id": 1, + "name": "Berry Brooks" + }, + { + "id": 2, + "name": "Griffin Larsen" + } + ], + "greeting": "Hello, Keller Valenzuela! You have 9 unread messages.", + "favoriteFruit": "strawberry" + }, + { + "_id": "58fa646dd6acde62d99d5283", + "index": 70, + "guid": "cdc093a8-acfe-4f5b-bbf4-426e17f6b07f", + "isActive": true, + "balance": "$3,972.53", + "picture": "http://placehold.it/32x32", + "age": 34, + "eyeColor": "green", + "name": "Shanna Gay", + "gender": "female", + "company": "GLOBOIL", + "email": "shannagay@globoil.com", + "phone": "+1 (825) 456-2279", + "address": "315 Caton Avenue, Groveville, Nebraska, 5888", + "registered": "2015-10-21T07:45:51 -02:00", + "latitude": 11.496581, + "longitude": 92.003771, + "tags": [ + "commodo", + "aliqua", + "dolore", + "nostrud", + "ipsum", + "officia", + "proident" + ], + "friends": [ + { + "id": 0, + "name": "Alexander Glover" + }, + { + "id": 1, + "name": "Bradley Randolph" + }, + { + "id": 2, + "name": "Warner Finch" + } + ], + "greeting": "Hello, Shanna Gay! You have 10 unread messages.", + "favoriteFruit": "banana" + }, + { + "_id": "58fa646d57d1d955ad3133b7", + "index": 71, + "guid": "b301e156-9d99-445e-8178-5ddd8f77f64e", + "isActive": false, + "balance": "$1,932.98", + "picture": "http://placehold.it/32x32", + "age": 39, + "eyeColor": "brown", + "name": "Chandler Carroll", + "gender": "male", + "company": "CANDECOR", + "email": "chandlercarroll@candecor.com", + "phone": "+1 (922) 585-2757", + "address": "544 Pulaski Street, Gilgo, Nevada, 2781", + "registered": "2016-03-24T07:02:16 -01:00", + "latitude": 89.136705, + "longitude": -78.050351, + "tags": [ + "in", + "occaecat", + "aliqua", + "minim", + "ut", + "ex", + "sit" + ], + "friends": [ + { + "id": 0, + "name": "Ida Erickson" + }, + { + "id": 1, + "name": "Coleen Peters" + }, + { + "id": 2, + "name": "Barron Shields" + } + ], + "greeting": "Hello, Chandler Carroll! You have 3 unread messages.", + "favoriteFruit": "banana" + }, + { + "_id": "58fa646d0a2271bb41cc739f", + "index": 72, + "guid": "c0b15c07-e224-4302-a8c0-df82447c7114", + "isActive": false, + "balance": "$2,527.83", + "picture": "http://placehold.it/32x32", + "age": 23, + "eyeColor": "brown", + "name": "Mcclain Herrera", + "gender": "male", + "company": "EYERIS", + "email": "mcclainherrera@eyeris.com", + "phone": "+1 (830) 498-2923", + "address": "332 Meserole Avenue, Drytown, Mississippi, 6200", + "registered": "2015-01-25T06:41:59 -01:00", + "latitude": 70.688802, + "longitude": 29.048923, + "tags": [ + "fugiat", + "quis", + "quis", + "minim", + "est", + "et", + "esse" + ], + "friends": [ + { + "id": 0, + "name": "Ruthie Buckley" + }, + { + "id": 1, + "name": "Rollins Curtis" + }, + { + "id": 2, + "name": "Lorna Key" + } + ], + "greeting": "Hello, Mcclain Herrera! You have 8 unread messages.", + "favoriteFruit": "banana" + }, + { + "_id": "58fa646d81d1b7d679650e03", + "index": 73, + "guid": "794c3ed6-bcc3-4173-b66a-8142b05bef3f", + "isActive": false, + "balance": "$1,110.74", + "picture": "http://placehold.it/32x32", + "age": 39, + "eyeColor": "green", + "name": "Myrna Mathis", + "gender": "female", + "company": "ZENTHALL", + "email": "myrnamathis@zenthall.com", + "phone": "+1 (993) 500-2210", + "address": "862 Halsey Street, Foscoe, New York, 2877", + "registered": "2015-01-01T07:49:57 -01:00", + "latitude": -58.371217, + "longitude": -127.79004, + "tags": [ + "duis", + "veniam", + "labore", + "amet", + "dolore", + "irure", + "tempor" + ], + "friends": [ + { + "id": 0, + "name": "Morgan Daugherty" + }, + { + "id": 1, + "name": "Wilkins Holmes" + }, + { + "id": 2, + "name": "Ella Patterson" + } + ], + "greeting": "Hello, Myrna Mathis! You have 6 unread messages.", + "favoriteFruit": "strawberry" + }, + { + "_id": "58fa646dc05ee21f4b7724f1", + "index": 74, + "guid": "56faf663-b97f-4c41-923c-924e056c6882", + "isActive": false, + "balance": "$1,453.59", + "picture": "http://placehold.it/32x32", + "age": 23, + "eyeColor": "blue", + "name": "Teri Finley", + "gender": "female", + "company": "ZILPHUR", + "email": "terifinley@zilphur.com", + "phone": "+1 (908) 582-2539", + "address": "461 Quentin Road, Wheaton, Oklahoma, 6775", + "registered": "2015-01-10T10:05:10 -01:00", + "latitude": -47.354044, + "longitude": 34.61937, + "tags": [ + "enim", + "nulla", + "ullamco", + "irure", + "nostrud", + "cupidatat", + "dolor" + ], + "friends": [ + { + "id": 0, + "name": "Boone Emerson" + }, + { + "id": 1, + "name": "Pennington Baxter" + }, + { + "id": 2, + "name": "Miranda Sweeney" + } + ], + "greeting": "Hello, Teri Finley! You have 7 unread messages.", + "favoriteFruit": "banana" + }, + { + "_id": "58fa646d6ca83df640f89e92", + "index": 75, + "guid": "285e0974-d51e-482e-90ed-6cb75c56e43f", + "isActive": false, + "balance": "$1,709.68", + "picture": "http://placehold.it/32x32", + "age": 30, + "eyeColor": "brown", + "name": "Stein Bean", + "gender": "male", + "company": "PEARLESEX", + "email": "steinbean@pearlesex.com", + "phone": "+1 (915) 407-2551", + "address": "527 Plymouth Street, Loveland, Maryland, 3731", + "registered": "2014-11-05T09:54:08 -01:00", + "latitude": 36.488153, + "longitude": -69.04741, + "tags": [ + "id", + "in", + "voluptate", + "cupidatat", + "eiusmod", + "do", + "consequat" + ], + "friends": [ + { + "id": 0, + "name": "Barbra Dorsey" + }, + { + "id": 1, + "name": "Audrey Farmer" + }, + { + "id": 2, + "name": "Liliana Barker" + } + ], + "greeting": "Hello, Stein Bean! You have 5 unread messages.", + "favoriteFruit": "strawberry" + }, + { + "_id": "58fa646d1b17a06de9b76ab4", + "index": 76, + "guid": "1e43d3dc-802e-4064-a504-2f3274510c96", + "isActive": false, + "balance": "$1,407.58", + "picture": "http://placehold.it/32x32", + "age": 22, + "eyeColor": "blue", + "name": "Morin Morse", + "gender": "male", + "company": "MUSIX", + "email": "morinmorse@musix.com", + "phone": "+1 (886) 417-3779", + "address": "416 Macdougal Street, Gulf, Wyoming, 3967", + "registered": "2015-09-23T01:19:46 -02:00", + "latitude": -37.559365, + "longitude": -128.533029, + "tags": [ + "sunt", + "dolor", + "velit", + "irure", + "occaecat", + "reprehenderit", + "ex" + ], + "friends": [ + { + "id": 0, + "name": "Brady Gaines" + }, + { + "id": 1, + "name": "Powers Irwin" + }, + { + "id": 2, + "name": "Lyons Leonard" + } + ], + "greeting": "Hello, Morin Morse! You have 4 unread messages.", + "favoriteFruit": "strawberry" + }, + { + "_id": "58fa646ddfa89531a925f1b6", + "index": 77, + "guid": "41646507-048d-4bec-ac36-026e97777267", + "isActive": false, + "balance": "$1,308.52", + "picture": "http://placehold.it/32x32", + "age": 38, + "eyeColor": "blue", + "name": "Malinda Jarvis", + "gender": "female", + "company": "OATFARM", + "email": "malindajarvis@oatfarm.com", + "phone": "+1 (962) 453-2047", + "address": "941 Colonial Road, Nord, Arizona, 4955", + "registered": "2015-03-30T04:46:18 -02:00", + "latitude": -77.024224, + "longitude": -7.764407, + "tags": [ + "voluptate", + "ipsum", + "aute", + "nulla", + "culpa", + "tempor", + "non" + ], + "friends": [ + { + "id": 0, + "name": "Lynette Preston" + }, + { + "id": 1, + "name": "Shepherd Riley" + }, + { + "id": 2, + "name": "Gray Mcguire" + } + ], + "greeting": "Hello, Malinda Jarvis! You have 5 unread messages.", + "favoriteFruit": "apple" + }, + { + "_id": "58fa646dfb5129928cac82c4", + "index": 78, + "guid": "9c565221-1528-4bd9-acc0-1860188f5103", + "isActive": false, + "balance": "$1,611.68", + "picture": "http://placehold.it/32x32", + "age": 35, + "eyeColor": "blue", + "name": "Soto Garza", + "gender": "male", + "company": "HIVEDOM", + "email": "sotogarza@hivedom.com", + "phone": "+1 (977) 528-2258", + "address": "673 Madison Street, Fivepointville, Texas, 7604", + "registered": "2016-10-24T07:08:28 -02:00", + "latitude": -72.614931, + "longitude": -3.073502, + "tags": [ + "quis", + "nulla", + "occaecat", + "consequat", + "consequat", + "est", + "laborum" + ], + "friends": [ + { + "id": 0, + "name": "Ward Travis" + }, + { + "id": 1, + "name": "Patel Fox" + }, + { + "id": 2, + "name": "Hodge Dominguez" + } + ], + "greeting": "Hello, Soto Garza! You have 2 unread messages.", + "favoriteFruit": "apple" + }, + { + "_id": "58fa646de3c73b55be3b2392", + "index": 79, + "guid": "5fc8084c-7b33-4b71-b56a-155ee73bf4d5", + "isActive": true, + "balance": "$1,193.72", + "picture": "http://placehold.it/32x32", + "age": 24, + "eyeColor": "blue", + "name": "Valeria Noble", + "gender": "female", + "company": "PETIGEMS", + "email": "valerianoble@petigems.com", + "phone": "+1 (968) 449-2505", + "address": "557 Clymer Street, Roy, Connecticut, 9406", + "registered": "2014-01-09T04:32:14 -01:00", + "latitude": 73.833128, + "longitude": -2.938819, + "tags": [ + "commodo", + "qui", + "non", + "excepteur", + "anim", + "enim", + "fugiat" + ], + "friends": [ + { + "id": 0, + "name": "Acevedo Boyle" + }, + { + "id": 1, + "name": "Christine Horn" + }, + { + "id": 2, + "name": "Madeleine Morris" + } + ], + "greeting": "Hello, Valeria Noble! You have 1 unread messages.", + "favoriteFruit": "apple" + }, + { + "_id": "58fa646d419298c39c8e65a3", + "index": 80, + "guid": "3feeeac2-1bc6-49b1-9504-98b9c811477b", + "isActive": true, + "balance": "$3,951.53", + "picture": "http://placehold.it/32x32", + "age": 37, + "eyeColor": "green", + "name": "Fisher Potts", + "gender": "male", + "company": "GEEKOLOGY", + "email": "fisherpotts@geekology.com", + "phone": "+1 (949) 552-2479", + "address": "912 Forest Place, Camas, Maine, 3673", + "registered": "2016-10-31T06:16:55 -01:00", + "latitude": 9.197336, + "longitude": -161.018046, + "tags": [ + "ea", + "in", + "incididunt", + "labore", + "occaecat", + "ea", + "sunt" + ], + "friends": [ + { + "id": 0, + "name": "Decker Larson" + }, + { + "id": 1, + "name": "Georgette Richardson" + }, + { + "id": 2, + "name": "Pearson Middleton" + } + ], + "greeting": "Hello, Fisher Potts! You have 10 unread messages.", + "favoriteFruit": "apple" + }, + { + "_id": "58fa646d492daca29b3ba028", + "index": 81, + "guid": "ed7f71be-b43d-4412-89d6-20c3bd5e1b8b", + "isActive": false, + "balance": "$2,966.53", + "picture": "http://placehold.it/32x32", + "age": 40, + "eyeColor": "green", + "name": "Merrill Beasley", + "gender": "male", + "company": "SOPRANO", + "email": "merrillbeasley@soprano.com", + "phone": "+1 (814) 485-2555", + "address": "227 Ford Street, Topanga, Idaho, 6526", + "registered": "2016-12-05T03:13:04 -01:00", + "latitude": -58.024144, + "longitude": -85.45299, + "tags": [ + "Lorem", + "irure", + "excepteur", + "non", + "mollit", + "exercitation", + "exercitation" + ], + "friends": [ + { + "id": 0, + "name": "Beth Wilkinson" + }, + { + "id": 1, + "name": "Elsie Black" + }, + { + "id": 2, + "name": "Dollie Kelley" + } + ], + "greeting": "Hello, Merrill Beasley! You have 8 unread messages.", + "favoriteFruit": "banana" + }, + { + "_id": "58fa646d927649a71e9025db", + "index": 82, + "guid": "39d9ea70-dd3b-4c05-b85b-959397ec8347", + "isActive": false, + "balance": "$2,284.04", + "picture": "http://placehold.it/32x32", + "age": 24, + "eyeColor": "green", + "name": "Dixon Small", + "gender": "male", + "company": "ISOTERNIA", + "email": "dixonsmall@isoternia.com", + "phone": "+1 (957) 494-2734", + "address": "986 Forbell Street, Trexlertown, Hawaii, 9149", + "registered": "2017-01-20T02:34:17 -01:00", + "latitude": -89.58325, + "longitude": -83.752222, + "tags": [ + "pariatur", + "ipsum", + "ut", + "dolor", + "id", + "ullamco", + "anim" + ], + "friends": [ + { + "id": 0, + "name": "Summer Walker" + }, + { + "id": 1, + "name": "Virginia Morrison" + }, + { + "id": 2, + "name": "Mona Edwards" + } + ], + "greeting": "Hello, Dixon Small! You have 2 unread messages.", + "favoriteFruit": "apple" + }, + { + "_id": "58fa646d2f976042d79a7146", + "index": 83, + "guid": "4c1374ad-3e53-4d70-acf0-a323c835f35f", + "isActive": true, + "balance": "$1,397.97", + "picture": "http://placehold.it/32x32", + "age": 36, + "eyeColor": "brown", + "name": "Roy Porter", + "gender": "male", + "company": "UNDERTAP", + "email": "royporter@undertap.com", + "phone": "+1 (803) 491-3075", + "address": "768 John Street, Coleville, Alaska, 1437", + "registered": "2017-02-25T10:31:16 -01:00", + "latitude": -14.493062, + "longitude": 83.564408, + "tags": [ + "eiusmod", + "veniam", + "ipsum", + "sint", + "culpa", + "labore", + "velit" + ], + "friends": [ + { + "id": 0, + "name": "Faith Salazar" + }, + { + "id": 1, + "name": "Kristin Haynes" + }, + { + "id": 2, + "name": "Celina Woods" + } + ], + "greeting": "Hello, Roy Porter! You have 9 unread messages.", + "favoriteFruit": "strawberry" + }, + { + "_id": "58fa646da60ff7cf82d78866", + "index": 84, + "guid": "edfa573c-6386-4494-82a1-006191f03c7d", + "isActive": true, + "balance": "$3,843.02", + "picture": "http://placehold.it/32x32", + "age": 37, + "eyeColor": "green", + "name": "Moore Bates", + "gender": "male", + "company": "KROG", + "email": "moorebates@krog.com", + "phone": "+1 (919) 573-4000", + "address": "473 Miller Avenue, Rew, Iowa, 5386", + "registered": "2016-08-14T06:29:59 -02:00", + "latitude": -3.561713, + "longitude": 35.342721, + "tags": [ + "pariatur", + "nisi", + "do", + "ullamco", + "consequat", + "proident", + "labore" + ], + "friends": [ + { + "id": 0, + "name": "James Curry" + }, + { + "id": 1, + "name": "Vonda Schwartz" + }, + { + "id": 2, + "name": "Shields Shepard" + } + ], + "greeting": "Hello, Moore Bates! You have 8 unread messages.", + "favoriteFruit": "strawberry" + }, + { + "_id": "58fa646d0ac925b65c2325ed", + "index": 85, + "guid": "29c5c1bd-0dd2-4295-a967-0b0d27056faf", + "isActive": true, + "balance": "$3,580.48", + "picture": "http://placehold.it/32x32", + "age": 27, + "eyeColor": "green", + "name": "Emily Stevens", + "gender": "female", + "company": "ELECTONIC", + "email": "emilystevens@electonic.com", + "phone": "+1 (940) 563-2789", + "address": "615 Goodwin Place, Cashtown, Michigan, 7501", + "registered": "2014-05-11T10:32:41 -02:00", + "latitude": 73.085973, + "longitude": -81.515059, + "tags": [ + "magna", + "ut", + "occaecat", + "aute", + "culpa", + "dolor", + "proident" + ], + "friends": [ + { + "id": 0, + "name": "Francis Walter" + }, + { + "id": 1, + "name": "Hopper Blair" + }, + { + "id": 2, + "name": "Jordan Stanton" + } + ], + "greeting": "Hello, Emily Stevens! You have 8 unread messages.", + "favoriteFruit": "banana" + }, + { + "_id": "58fa646dc12ad81e55f05bbe", + "index": 86, + "guid": "fd3cc030-b17a-4307-bd80-affe95ee8bb8", + "isActive": false, + "balance": "$3,613.75", + "picture": "http://placehold.it/32x32", + "age": 28, + "eyeColor": "green", + "name": "Kirby Powell", + "gender": "male", + "company": "GRACKER", + "email": "kirbypowell@gracker.com", + "phone": "+1 (879) 424-3041", + "address": "648 Alabama Avenue, Wyano, North Dakota, 3436", + "registered": "2014-08-06T09:33:36 -02:00", + "latitude": 63.416412, + "longitude": 85.705391, + "tags": [ + "velit", + "enim", + "pariatur", + "dolore", + "deserunt", + "Lorem", + "aliquip" + ], + "friends": [ + { + "id": 0, + "name": "Dunn Conner" + }, + { + "id": 1, + "name": "Hernandez Flores" + }, + { + "id": 2, + "name": "Sybil Soto" + } + ], + "greeting": "Hello, Kirby Powell! You have 5 unread messages.", + "favoriteFruit": "strawberry" + }, + { + "_id": "58fa646de35bffff823ac35c", + "index": 87, + "guid": "7dd312de-dbc0-46c6-ba07-9ec0df483755", + "isActive": true, + "balance": "$3,146.50", + "picture": "http://placehold.it/32x32", + "age": 26, + "eyeColor": "brown", + "name": "Alison Kramer", + "gender": "female", + "company": "WEBIOTIC", + "email": "alisonkramer@webiotic.com", + "phone": "+1 (936) 531-2953", + "address": "720 Kingston Avenue, Cutter, Guam, 3910", + "registered": "2017-01-31T04:20:17 -01:00", + "latitude": -19.927529, + "longitude": 25.18981, + "tags": [ + "elit", + "tempor", + "officia", + "ut", + "do", + "cupidatat", + "dolor" + ], + "friends": [ + { + "id": 0, + "name": "Bird Fernandez" + }, + { + "id": 1, + "name": "Walls Bruce" + }, + { + "id": 2, + "name": "Fulton Drake" + } + ], + "greeting": "Hello, Alison Kramer! You have 7 unread messages.", + "favoriteFruit": "strawberry" + }, + { + "_id": "58fa646d73d27a4b03bb80ee", + "index": 88, + "guid": "bddaf610-0187-409f-aec4-5ba86d1459e1", + "isActive": false, + "balance": "$2,949.26", + "picture": "http://placehold.it/32x32", + "age": 21, + "eyeColor": "brown", + "name": "Lizzie Morrow", + "gender": "female", + "company": "PAPRICUT", + "email": "lizziemorrow@papricut.com", + "phone": "+1 (987) 599-3391", + "address": "191 Blake Avenue, Savannah, Virginia, 7252", + "registered": "2014-08-19T09:05:32 -02:00", + "latitude": 36.346384, + "longitude": -109.208162, + "tags": [ + "nulla", + "quis", + "irure", + "quis", + "laboris", + "Lorem", + "ipsum" + ], + "friends": [ + { + "id": 0, + "name": "Fernandez Stark" + }, + { + "id": 1, + "name": "Robles Crawford" + }, + { + "id": 2, + "name": "Rosemary Todd" + } + ], + "greeting": "Hello, Lizzie Morrow! You have 7 unread messages.", + "favoriteFruit": "banana" + }, + { + "_id": "58fa646da907f7d9e7bf8914", + "index": 89, + "guid": "92320867-3b54-4ab3-acd8-5a80e4b4d9c1", + "isActive": false, + "balance": "$3,492.69", + "picture": "http://placehold.it/32x32", + "age": 37, + "eyeColor": "brown", + "name": "Darlene Moore", + "gender": "female", + "company": "KOFFEE", + "email": "darlenemoore@koffee.com", + "phone": "+1 (887) 502-2164", + "address": "490 Catherine Street, Springhill, West Virginia, 7515", + "registered": "2014-02-20T09:17:04 -01:00", + "latitude": -13.214116, + "longitude": 131.231308, + "tags": [ + "consectetur", + "ut", + "laborum", + "duis", + "adipisicing", + "velit", + "exercitation" + ], + "friends": [ + { + "id": 0, + "name": "Lynn Ratliff" + }, + { + "id": 1, + "name": "Melendez Bennett" + }, + { + "id": 2, + "name": "Delia Atkins" + } + ], + "greeting": "Hello, Darlene Moore! You have 5 unread messages.", + "favoriteFruit": "strawberry" + }, + { + "_id": "58fa646d284ff51b946c9346", + "index": 90, + "guid": "7c83cad9-f51a-4e7e-af59-881e584d97d2", + "isActive": true, + "balance": "$2,202.45", + "picture": "http://placehold.it/32x32", + "age": 28, + "eyeColor": "green", + "name": "Consuelo Patel", + "gender": "female", + "company": "BILLMED", + "email": "consuelopatel@billmed.com", + "phone": "+1 (959) 420-2794", + "address": "924 Tiffany Place, Bartonsville, Kentucky, 9968", + "registered": "2016-03-31T09:23:28 -02:00", + "latitude": 55.703127, + "longitude": 54.070006, + "tags": [ + "id", + "incididunt", + "incididunt", + "amet", + "anim", + "dolor", + "ipsum" + ], + "friends": [ + { + "id": 0, + "name": "Helen Stewart" + }, + { + "id": 1, + "name": "Kramer Ramos" + }, + { + "id": 2, + "name": "Alba Daniels" + } + ], + "greeting": "Hello, Consuelo Patel! You have 9 unread messages.", + "favoriteFruit": "apple" + }, + { + "_id": "58fa646d89e96b74d8ca7854", + "index": 91, + "guid": "54b09276-2ba7-4c41-b69d-856c87e9506c", + "isActive": true, + "balance": "$3,810.19", + "picture": "http://placehold.it/32x32", + "age": 20, + "eyeColor": "green", + "name": "Compton Cameron", + "gender": "male", + "company": "ACCUPHARM", + "email": "comptoncameron@accupharm.com", + "phone": "+1 (850) 405-3734", + "address": "581 Hoyts Lane, Eagletown, Northern Mariana Islands, 1515", + "registered": "2015-06-15T07:47:59 -02:00", + "latitude": -82.414119, + "longitude": 58.652098, + "tags": [ + "cupidatat", + "aliquip", + "ea", + "qui", + "et", + "elit", + "nisi" + ], + "friends": [ + { + "id": 0, + "name": "Carol Spence" + }, + { + "id": 1, + "name": "Beach Oconnor" + }, + { + "id": 2, + "name": "Clare Zamora" + } + ], + "greeting": "Hello, Compton Cameron! You have 8 unread messages.", + "favoriteFruit": "strawberry" + }, + { + "_id": "58fa646dc3faec5f1c58e747", + "index": 92, + "guid": "3e3ed848-9ce0-4dec-9ca0-f9281578a35d", + "isActive": true, + "balance": "$2,631.80", + "picture": "http://placehold.it/32x32", + "age": 37, + "eyeColor": "green", + "name": "Sandoval King", + "gender": "male", + "company": "ZISIS", + "email": "sandovalking@zisis.com", + "phone": "+1 (922) 412-2186", + "address": "645 Bay Street, Lacomb, Virgin Islands, 2069", + "registered": "2014-06-25T05:57:37 -02:00", + "latitude": -49.124653, + "longitude": 4.474034, + "tags": [ + "duis", + "aute", + "consequat", + "id", + "laborum", + "officia", + "eiusmod" + ], + "friends": [ + { + "id": 0, + "name": "Karen Tucker" + }, + { + "id": 1, + "name": "White Bryant" + }, + { + "id": 2, + "name": "Turner Griffith" + } + ], + "greeting": "Hello, Sandoval King! You have 7 unread messages.", + "favoriteFruit": "banana" + }, + { + "_id": "58fa646d77b404a27884587f", + "index": 93, + "guid": "e4a14bf5-7bf5-4326-9168-8e5ddec4e950", + "isActive": false, + "balance": "$3,477.42", + "picture": "http://placehold.it/32x32", + "age": 23, + "eyeColor": "brown", + "name": "Glenn Ferguson", + "gender": "male", + "company": "ZILIDIUM", + "email": "glennferguson@zilidium.com", + "phone": "+1 (954) 521-2233", + "address": "156 Wythe Place, Fairacres, Federated States Of Micronesia, 1557", + "registered": "2014-10-29T03:16:45 -01:00", + "latitude": -32.281968, + "longitude": 132.755574, + "tags": [ + "anim", + "excepteur", + "eu", + "nisi", + "occaecat", + "nostrud", + "eiusmod" + ], + "friends": [ + { + "id": 0, + "name": "Richards Paul" + }, + { + "id": 1, + "name": "Flores Pearson" + }, + { + "id": 2, + "name": "Nash Owens" + } + ], + "greeting": "Hello, Glenn Ferguson! You have 3 unread messages.", + "favoriteFruit": "strawberry" + }, + { + "_id": "58fa646d68304e1ff7114424", + "index": 94, + "guid": "85ff3a1e-b7d2-4a21-b8b0-9b43abd0eca2", + "isActive": true, + "balance": "$2,401.89", + "picture": "http://placehold.it/32x32", + "age": 22, + "eyeColor": "blue", + "name": "Adeline Joyner", + "gender": "female", + "company": "SCENTRIC", + "email": "adelinejoyner@scentric.com", + "phone": "+1 (955) 538-2127", + "address": "549 Grand Street, Canterwood, Utah, 3548", + "registered": "2014-06-14T09:57:15 -02:00", + "latitude": 30.146503, + "longitude": -47.223023, + "tags": [ + "voluptate", + "officia", + "non", + "ex", + "duis", + "et", + "minim" + ], + "friends": [ + { + "id": 0, + "name": "Drake Hale" + }, + { + "id": 1, + "name": "Christensen Armstrong" + }, + { + "id": 2, + "name": "Hanson Watkins" + } + ], + "greeting": "Hello, Adeline Joyner! You have 4 unread messages.", + "favoriteFruit": "banana" + }, + { + "_id": "58fa646d8793da93732a6eb9", + "index": 95, + "guid": "3587b9d3-cda3-4001-925a-87d6877ef4c7", + "isActive": true, + "balance": "$3,545.34", + "picture": "http://placehold.it/32x32", + "age": 23, + "eyeColor": "green", + "name": "Steele Cohen", + "gender": "male", + "company": "VERTIDE", + "email": "steelecohen@vertide.com", + "phone": "+1 (896) 444-2215", + "address": "216 Harden Street, Rowe, Georgia, 3268", + "registered": "2016-01-24T12:19:54 -01:00", + "latitude": -80.65651, + "longitude": 66.542419, + "tags": [ + "adipisicing", + "elit", + "consequat", + "exercitation", + "sunt", + "amet", + "laborum" + ], + "friends": [ + { + "id": 0, + "name": "Lambert Knapp" + }, + { + "id": 1, + "name": "Georgina Richmond" + }, + { + "id": 2, + "name": "Hampton Calderon" + } + ], + "greeting": "Hello, Steele Cohen! You have 6 unread messages.", + "favoriteFruit": "strawberry" + }, + { + "_id": "58fa646de7a01dd2ba1e3bd6", + "index": 96, + "guid": "d356f69e-995c-4fe5-a831-c349c6030342", + "isActive": true, + "balance": "$2,947.31", + "picture": "http://placehold.it/32x32", + "age": 21, + "eyeColor": "green", + "name": "Cantu Cantrell", + "gender": "male", + "company": "JIMBIES", + "email": "cantucantrell@jimbies.com", + "phone": "+1 (931) 475-3741", + "address": "506 Bay Parkway, Colton, Pennsylvania, 2474", + "registered": "2016-04-20T10:27:10 -02:00", + "latitude": -76.295054, + "longitude": 63.991426, + "tags": [ + "nulla", + "est", + "id", + "dolore", + "consectetur", + "dolore", + "irure" + ], + "friends": [ + { + "id": 0, + "name": "Mercado Munoz" + }, + { + "id": 1, + "name": "Eula Vaughn" + }, + { + "id": 2, + "name": "Antoinette Church" + } + ], + "greeting": "Hello, Cantu Cantrell! You have 5 unread messages.", + "favoriteFruit": "apple" + }, + { + "_id": "58fa646d3c6e1f7675cfac0f", + "index": 97, + "guid": "38f9ead3-0bd3-49d7-ba6a-85f6007734ba", + "isActive": true, + "balance": "$1,224.06", + "picture": "http://placehold.it/32x32", + "age": 39, + "eyeColor": "green", + "name": "Hudson Butler", + "gender": "male", + "company": "PHEAST", + "email": "hudsonbutler@pheast.com", + "phone": "+1 (908) 595-2011", + "address": "455 Sullivan Street, Golconda, Indiana, 3642", + "registered": "2014-06-20T04:02:28 -02:00", + "latitude": 85.315515, + "longitude": 27.32936, + "tags": [ + "aliqua", + "cillum", + "non", + "Lorem", + "qui", + "consequat", + "culpa" + ], + "friends": [ + { + "id": 0, + "name": "Harrison Robertson" + }, + { + "id": 1, + "name": "Mitchell Wyatt" + }, + { + "id": 2, + "name": "Dyer Pena" + } + ], + "greeting": "Hello, Hudson Butler! You have 4 unread messages.", + "favoriteFruit": "strawberry" + }, + { + "_id": "58fa646da0bf70d54c620dc1", + "index": 98, + "guid": "32394dc0-3290-4e2f-a0b4-c93d6166e3ee", + "isActive": true, + "balance": "$3,092.67", + "picture": "http://placehold.it/32x32", + "age": 28, + "eyeColor": "green", + "name": "Byers Johnson", + "gender": "male", + "company": "PIGZART", + "email": "byersjohnson@pigzart.com", + "phone": "+1 (979) 452-2962", + "address": "220 Fanchon Place, Coinjock, District Of Columbia, 2867", + "registered": "2014-09-26T12:03:34 -02:00", + "latitude": -22.284103, + "longitude": 0.759555, + "tags": [ + "ipsum", + "veniam", + "ea", + "dolore", + "fugiat", + "culpa", + "id" + ], + "friends": [ + { + "id": 0, + "name": "Neva Garcia" + }, + { + "id": 1, + "name": "Amalia Harrington" + }, + { + "id": 2, + "name": "Cain Ryan" + } + ], + "greeting": "Hello, Byers Johnson! You have 4 unread messages.", + "favoriteFruit": "apple" + }, + { + "_id": "58fa646df37896247bcaac9f", + "index": 99, + "guid": "7ce04f36-22b2-4bd0-9a33-d002bc6356ba", + "isActive": true, + "balance": "$2,223.41", + "picture": "http://placehold.it/32x32", + "age": 23, + "eyeColor": "brown", + "name": "Wynn Goodman", + "gender": "male", + "company": "EZENTIA", + "email": "wynngoodman@ezentia.com", + "phone": "+1 (954) 433-3278", + "address": "627 Hastings Street, Williston, South Dakota, 594", + "registered": "2015-11-27T10:23:55 -01:00", + "latitude": -22.494524, + "longitude": 40.256038, + "tags": [ + "sint", + "ea", + "nisi", + "ipsum", + "sint", + "quis", + "anim" + ], + "friends": [ + { + "id": 0, + "name": "Newton Flynn" + }, + { + "id": 1, + "name": "Cannon Maynard" + }, + { + "id": 2, + "name": "Bettye Avila" + } + ], + "greeting": "Hello, Wynn Goodman! You have 2 unread messages.", + "favoriteFruit": "apple" + }, + { + "_id": "58fa646daedc8b1bb75af807", + "index": 100, + "guid": "100c1a9e-02e2-469e-b707-8710b3137f75", + "isActive": true, + "balance": "$2,986.52", + "picture": "http://placehold.it/32x32", + "age": 35, + "eyeColor": "brown", + "name": "Knowles Rocha", + "gender": "male", + "company": "ACUMENTOR", + "email": "knowlesrocha@acumentor.com", + "phone": "+1 (804) 538-3561", + "address": "779 Nevins Street, Brandywine, Ohio, 505", + "registered": "2014-01-26T08:55:03 -01:00", + "latitude": 5.387132, + "longitude": 69.204198, + "tags": [ + "cillum", + "pariatur", + "irure", + "ad", + "nulla", + "eiusmod", + "esse" + ], + "friends": [ + { + "id": 0, + "name": "Mcpherson Atkinson" + }, + { + "id": 1, + "name": "Maura Compton" + }, + { + "id": 2, + "name": "Justine Nguyen" + } + ], + "greeting": "Hello, Knowles Rocha! You have 4 unread messages.", + "favoriteFruit": "banana" + }, + { + "_id": "58fa646d4cea40b1ea6a5593", + "index": 101, + "guid": "ea3a446f-1edd-4d6c-b92c-4e5446fb38e3", + "isActive": false, + "balance": "$1,630.01", + "picture": "http://placehold.it/32x32", + "age": 32, + "eyeColor": "blue", + "name": "Kaitlin Huber", + "gender": "female", + "company": "CORIANDER", + "email": "kaitlinhuber@coriander.com", + "phone": "+1 (962) 479-2133", + "address": "584 Friel Place, Yorklyn, Palau, 4221", + "registered": "2016-06-12T08:39:58 -02:00", + "latitude": -37.20637, + "longitude": -99.426599, + "tags": [ + "est", + "eu", + "adipisicing", + "laborum", + "quis", + "tempor", + "sit" + ], + "friends": [ + { + "id": 0, + "name": "Parrish Banks" + }, + { + "id": 1, + "name": "Fry Nichols" + }, + { + "id": 2, + "name": "Eve Boyd" + } + ], + "greeting": "Hello, Kaitlin Huber! You have 1 unread messages.", + "favoriteFruit": "apple" + }, + { + "_id": "58fa646d67d4d5b34df78691", + "index": 102, + "guid": "259c0bb2-13ed-4274-bfb8-84bfd51ee77d", + "isActive": false, + "balance": "$1,801.52", + "picture": "http://placehold.it/32x32", + "age": 38, + "eyeColor": "green", + "name": "Eleanor Turner", + "gender": "female", + "company": "ISOSURE", + "email": "eleanorturner@isosure.com", + "phone": "+1 (906) 572-2993", + "address": "194 Metrotech Courtr, Montura, Louisiana, 6246", + "registered": "2015-06-20T11:55:50 -02:00", + "latitude": 81.620755, + "longitude": -95.421532, + "tags": [ + "enim", + "dolore", + "consequat", + "culpa", + "sint", + "non", + "dolor" + ], + "friends": [ + { + "id": 0, + "name": "Jocelyn Austin" + }, + { + "id": 1, + "name": "Roseann Miles" + }, + { + "id": 2, + "name": "Clarke Dickson" + } + ], + "greeting": "Hello, Eleanor Turner! You have 8 unread messages.", + "favoriteFruit": "apple" + }, + { + "_id": "58fa646dc3f0c86735d79a14", + "index": 103, + "guid": "b8151c1b-e461-4075-a86b-458c984d03b6", + "isActive": true, + "balance": "$2,546.74", + "picture": "http://placehold.it/32x32", + "age": 24, + "eyeColor": "brown", + "name": "Oneil Justice", + "gender": "male", + "company": "KAGE", + "email": "oneiljustice@kage.com", + "phone": "+1 (951) 442-2829", + "address": "920 Legion Street, Elfrida, Montana, 697", + "registered": "2015-06-29T04:12:54 -02:00", + "latitude": -13.095961, + "longitude": 10.327859, + "tags": [ + "anim", + "exercitation", + "Lorem", + "officia", + "sint", + "mollit", + "reprehenderit" + ], + "friends": [ + { + "id": 0, + "name": "Bauer Holloway" + }, + { + "id": 1, + "name": "Velasquez Trevino" + }, + { + "id": 2, + "name": "Ware Horton" + } + ], + "greeting": "Hello, Oneil Justice! You have 3 unread messages.", + "favoriteFruit": "apple" + }, + { + "_id": "58fa646df7fcb12af891dff4", + "index": 104, + "guid": "4305e081-0a38-47a8-bab0-7cfafca4e8fe", + "isActive": false, + "balance": "$3,563.58", + "picture": "http://placehold.it/32x32", + "age": 35, + "eyeColor": "brown", + "name": "Lenore Fitzpatrick", + "gender": "female", + "company": "VELITY", + "email": "lenorefitzpatrick@velity.com", + "phone": "+1 (952) 485-2274", + "address": "130 Canton Court, Alderpoint, Illinois, 2655", + "registered": "2014-05-31T06:48:05 -02:00", + "latitude": -38.271618, + "longitude": 111.481094, + "tags": [ + "Lorem", + "magna", + "consectetur", + "consequat", + "deserunt", + "sint", + "laborum" + ], + "friends": [ + { + "id": 0, + "name": "Cruz Hill" + }, + { + "id": 1, + "name": "Ilene Hogan" + }, + { + "id": 2, + "name": "Lucas Steele" + } + ], + "greeting": "Hello, Lenore Fitzpatrick! You have 8 unread messages.", + "favoriteFruit": "strawberry" + }, + { + "_id": "58fa646dbe0ac2394d03c231", + "index": 105, + "guid": "c0b97fdf-fe1f-407d-b5fc-5e2a12c6d3c5", + "isActive": false, + "balance": "$1,554.83", + "picture": "http://placehold.it/32x32", + "age": 37, + "eyeColor": "green", + "name": "Reese Roy", + "gender": "male", + "company": "ZILLAN", + "email": "reeseroy@zillan.com", + "phone": "+1 (834) 468-2731", + "address": "986 Eldert Street, Rushford, Minnesota, 9754", + "registered": "2016-01-26T08:33:06 -01:00", + "latitude": -72.698969, + "longitude": 118.627661, + "tags": [ + "voluptate", + "et", + "qui", + "veniam", + "adipisicing", + "aute", + "elit" + ], + "friends": [ + { + "id": 0, + "name": "Rojas Levine" + }, + { + "id": 1, + "name": "Gentry Valdez" + }, + { + "id": 2, + "name": "Gabrielle Burt" + } + ], + "greeting": "Hello, Reese Roy! You have 9 unread messages.", + "favoriteFruit": "banana" + }, + { + "_id": "58fa646d2f5f1a65b50f46ad", + "index": 106, + "guid": "18b440fc-86c7-4537-9466-5900b383697a", + "isActive": true, + "balance": "$1,841.19", + "picture": "http://placehold.it/32x32", + "age": 32, + "eyeColor": "green", + "name": "Roxanne Dale", + "gender": "female", + "company": "ORBIFLEX", + "email": "roxannedale@orbiflex.com", + "phone": "+1 (915) 536-3775", + "address": "872 Herkimer Place, Idamay, South Carolina, 7291", + "registered": "2016-10-23T03:39:50 -02:00", + "latitude": 25.592962, + "longitude": -24.120857, + "tags": [ + "qui", + "tempor", + "veniam", + "est", + "Lorem", + "ullamco", + "ipsum" + ], + "friends": [ + { + "id": 0, + "name": "Allen Mckinney" + }, + { + "id": 1, + "name": "Dianna Harding" + }, + { + "id": 2, + "name": "Barry Lara" + } + ], + "greeting": "Hello, Roxanne Dale! You have 7 unread messages.", + "favoriteFruit": "apple" + }, + { + "_id": "58fa646d6378cd5acfe67991", + "index": 107, + "guid": "55724354-4550-4d43-bdf0-cffcfb378d3b", + "isActive": false, + "balance": "$1,038.89", + "picture": "http://placehold.it/32x32", + "age": 35, + "eyeColor": "blue", + "name": "Vang Mclean", + "gender": "male", + "company": "NORSUL", + "email": "vangmclean@norsul.com", + "phone": "+1 (879) 429-3534", + "address": "231 Newel Street, Deputy, New Mexico, 9365", + "registered": "2014-11-03T11:52:29 -01:00", + "latitude": -64.834858, + "longitude": -63.923681, + "tags": [ + "excepteur", + "dolore", + "laboris", + "qui", + "occaecat", + "adipisicing", + "dolore" + ], + "friends": [ + { + "id": 0, + "name": "Rebekah Figueroa" + }, + { + "id": 1, + "name": "Terry Rivers" + }, + { + "id": 2, + "name": "Schmidt Dawson" + } + ], + "greeting": "Hello, Vang Mclean! You have 6 unread messages.", + "favoriteFruit": "banana" + }, + { + "_id": "58fa646db88b90c724a381ac", + "index": 108, + "guid": "0b340d43-bc54-4446-8ae9-bc8347470c71", + "isActive": false, + "balance": "$2,732.74", + "picture": "http://placehold.it/32x32", + "age": 20, + "eyeColor": "brown", + "name": "Bowers Gill", + "gender": "male", + "company": "ZILLADYNE", + "email": "bowersgill@zilladyne.com", + "phone": "+1 (852) 597-3675", + "address": "575 Nautilus Avenue, Zeba, North Carolina, 4815", + "registered": "2014-03-08T03:11:41 -01:00", + "latitude": 55.638622, + "longitude": 111.440553, + "tags": [ + "laboris", + "cupidatat", + "aliqua", + "excepteur", + "nisi", + "deserunt", + "ut" + ], + "friends": [ + { + "id": 0, + "name": "Wendi Reeves" + }, + { + "id": 1, + "name": "Stacy Mcclain" + }, + { + "id": 2, + "name": "Maryann Alston" + } + ], + "greeting": "Hello, Bowers Gill! You have 1 unread messages.", + "favoriteFruit": "banana" + }, + { + "_id": "58fa646d6a0c3dd9a8ba3646", + "index": 109, + "guid": "41d4096e-2c74-42e2-a7e8-babced25a2e2", + "isActive": true, + "balance": "$2,082.73", + "picture": "http://placehold.it/32x32", + "age": 33, + "eyeColor": "blue", + "name": "Cantrell Cooke", + "gender": "male", + "company": "CEPRENE", + "email": "cantrellcooke@ceprene.com", + "phone": "+1 (886) 525-3882", + "address": "580 Estate Road, Beyerville, Massachusetts, 4650", + "registered": "2016-07-22T08:27:10 -02:00", + "latitude": -77.498482, + "longitude": 173.606368, + "tags": [ + "non", + "ipsum", + "Lorem", + "aliquip", + "occaecat", + "do", + "anim" + ], + "friends": [ + { + "id": 0, + "name": "Lana Sexton" + }, + { + "id": 1, + "name": "Sonja Savage" + }, + { + "id": 2, + "name": "Day Williams" + } + ], + "greeting": "Hello, Cantrell Cooke! You have 2 unread messages.", + "favoriteFruit": "strawberry" + }, + { + "_id": "58fa646d4569174ff05d52b6", + "index": 110, + "guid": "f475d08b-d82f-438a-9bab-0ed58b707521", + "isActive": true, + "balance": "$1,388.88", + "picture": "http://placehold.it/32x32", + "age": 30, + "eyeColor": "green", + "name": "Maxwell Rivera", + "gender": "male", + "company": "DIGITALUS", + "email": "maxwellrivera@digitalus.com", + "phone": "+1 (895) 585-2040", + "address": "267 Danforth Street, Spokane, Kansas, 313", + "registered": "2014-06-26T10:05:34 -02:00", + "latitude": -40.38383, + "longitude": 128.725271, + "tags": [ + "ipsum", + "occaecat", + "occaecat", + "cupidatat", + "fugiat", + "labore", + "irure" + ], + "friends": [ + { + "id": 0, + "name": "Lidia Witt" + }, + { + "id": 1, + "name": "Flowers Reilly" + }, + { + "id": 2, + "name": "Lindsey Cotton" + } + ], + "greeting": "Hello, Maxwell Rivera! You have 1 unread messages.", + "favoriteFruit": "banana" + }, + { + "_id": "58fa646d1347ab945233f833", + "index": 111, + "guid": "b395fc62-4cab-49d0-9a96-d911ac010500", + "isActive": true, + "balance": "$2,019.57", + "picture": "http://placehold.it/32x32", + "age": 29, + "eyeColor": "green", + "name": "Lacey Mccray", + "gender": "female", + "company": "EMERGENT", + "email": "laceymccray@emergent.com", + "phone": "+1 (809) 482-2975", + "address": "332 Hunterfly Place, Farmers, American Samoa, 8441", + "registered": "2016-05-04T03:30:20 -02:00", + "latitude": 2.393145, + "longitude": -144.537217, + "tags": [ + "nulla", + "tempor", + "ut", + "pariatur", + "consectetur", + "ad", + "nulla" + ], + "friends": [ + { + "id": 0, + "name": "Kristie Barrett" + }, + { + "id": 1, + "name": "Carolina Logan" + }, + { + "id": 2, + "name": "Carter Brown" + } + ], + "greeting": "Hello, Lacey Mccray! You have 1 unread messages.", + "favoriteFruit": "apple" + }, + { + "_id": "58fa646dadb3df3925387ffc", + "index": 112, + "guid": "3a416f9b-8e5d-4d25-b977-01321148cdeb", + "isActive": true, + "balance": "$2,532.99", + "picture": "http://placehold.it/32x32", + "age": 36, + "eyeColor": "blue", + "name": "Wolfe Garrett", + "gender": "male", + "company": "BITREX", + "email": "wolfegarrett@bitrex.com", + "phone": "+1 (831) 430-3253", + "address": "995 Banner Avenue, Orason, Tennessee, 7857", + "registered": "2015-05-22T02:44:43 -02:00", + "latitude": -8.555745, + "longitude": 20.859709, + "tags": [ + "amet", + "consectetur", + "cillum", + "excepteur", + "enim", + "voluptate", + "nisi" + ], + "friends": [ + { + "id": 0, + "name": "Phillips Acosta" + }, + { + "id": 1, + "name": "Mathews Wilkins" + }, + { + "id": 2, + "name": "Lynda Graves" + } + ], + "greeting": "Hello, Wolfe Garrett! You have 2 unread messages.", + "favoriteFruit": "banana" + }, + { + "_id": "58fa646d42beccbd0edfacc6", + "index": 113, + "guid": "bacbc39f-7dbe-4ea0-905e-702f25b1661e", + "isActive": true, + "balance": "$2,895.09", + "picture": "http://placehold.it/32x32", + "age": 38, + "eyeColor": "green", + "name": "Marisol Weber", + "gender": "female", + "company": "DIGIFAD", + "email": "marisolweber@digifad.com", + "phone": "+1 (877) 456-2622", + "address": "496 Brightwater Avenue, Kenwood, Florida, 4838", + "registered": "2015-06-09T06:36:09 -02:00", + "latitude": -79.696655, + "longitude": 134.111019, + "tags": [ + "exercitation", + "pariatur", + "commodo", + "anim", + "sunt", + "id", + "minim" + ], + "friends": [ + { + "id": 0, + "name": "May Cross" + }, + { + "id": 1, + "name": "Finley Hays" + }, + { + "id": 2, + "name": "David Hanson" + } + ], + "greeting": "Hello, Marisol Weber! You have 10 unread messages.", + "favoriteFruit": "strawberry" + }, + { + "_id": "58fa646d43f81e2a8c6aaba1", + "index": 114, + "guid": "09d10965-77a6-4548-b75a-184c2203ee71", + "isActive": false, + "balance": "$3,959.21", + "picture": "http://placehold.it/32x32", + "age": 37, + "eyeColor": "blue", + "name": "Bernice Walsh", + "gender": "female", + "company": "MOTOVATE", + "email": "bernicewalsh@motovate.com", + "phone": "+1 (863) 514-3224", + "address": "383 Chapel Street, Northridge, Missouri, 4567", + "registered": "2014-12-07T08:20:04 -01:00", + "latitude": 57.67878, + "longitude": -84.800733, + "tags": [ + "tempor", + "dolore", + "nisi", + "labore", + "nostrud", + "aute", + "aliquip" + ], + "friends": [ + { + "id": 0, + "name": "Freida Ortiz" + }, + { + "id": 1, + "name": "Eva Goodwin" + }, + { + "id": 2, + "name": "Marisa Workman" + } + ], + "greeting": "Hello, Bernice Walsh! You have 1 unread messages.", + "favoriteFruit": "apple" + }, + { + "_id": "58fa646d9f1df21cab418959", + "index": 115, + "guid": "f4ebdcb3-f5db-4306-ba17-02828f686b03", + "isActive": false, + "balance": "$3,302.61", + "picture": "http://placehold.it/32x32", + "age": 38, + "eyeColor": "blue", + "name": "Deanna Carney", + "gender": "female", + "company": "EXOPLODE", + "email": "deannacarney@exoplode.com", + "phone": "+1 (956) 483-2731", + "address": "231 Manhattan Court, Elbert, Marshall Islands, 5073", + "registered": "2017-04-04T06:26:09 -02:00", + "latitude": -61.544889, + "longitude": -58.847775, + "tags": [ + "quis", + "deserunt", + "ullamco", + "ex", + "Lorem", + "fugiat", + "enim" + ], + "friends": [ + { + "id": 0, + "name": "Gilda Kirkland" + }, + { + "id": 1, + "name": "Mckinney Albert" + }, + { + "id": 2, + "name": "Hart Rosales" + } + ], + "greeting": "Hello, Deanna Carney! You have 2 unread messages.", + "favoriteFruit": "strawberry" + }, + { + "_id": "58fa646dd779da49e22b2767", + "index": 116, + "guid": "a1bb8537-5524-4374-b43c-61c58af58d6f", + "isActive": false, + "balance": "$3,236.44", + "picture": "http://placehold.it/32x32", + "age": 27, + "eyeColor": "green", + "name": "Snow Rios", + "gender": "male", + "company": "SPRINGBEE", + "email": "snowrios@springbee.com", + "phone": "+1 (896) 423-3893", + "address": "193 Burnett Street, Malo, Washington, 1537", + "registered": "2016-02-24T01:53:45 -01:00", + "latitude": -74.859544, + "longitude": -80.392155, + "tags": [ + "incididunt", + "do", + "dolor", + "eu", + "in", + "laborum", + "ex" + ], + "friends": [ + { + "id": 0, + "name": "Bartlett Acevedo" + }, + { + "id": 1, + "name": "Fletcher Richards" + }, + { + "id": 2, + "name": "Guerrero Yang" + } + ], + "greeting": "Hello, Snow Rios! You have 5 unread messages.", + "favoriteFruit": "strawberry" + }, + { + "_id": "58fa646da45fcaf101376acf", + "index": 117, + "guid": "2ae0876f-63e1-4e11-a85a-dd48e0f4890a", + "isActive": true, + "balance": "$3,624.01", + "picture": "http://placehold.it/32x32", + "age": 36, + "eyeColor": "brown", + "name": "Lelia Burns", + "gender": "female", + "company": "COMBOGEN", + "email": "leliaburns@combogen.com", + "phone": "+1 (914) 449-3856", + "address": "614 Bijou Avenue, Worcester, Colorado, 4953", + "registered": "2015-05-06T12:48:58 -02:00", + "latitude": 36.910016, + "longitude": -52.661459, + "tags": [ + "culpa", + "ea", + "veniam", + "ipsum", + "laborum", + "ipsum", + "culpa" + ], + "friends": [ + { + "id": 0, + "name": "Greene Manning" + }, + { + "id": 1, + "name": "Watkins Chang" + }, + { + "id": 2, + "name": "Barrera Rosa" + } + ], + "greeting": "Hello, Lelia Burns! You have 1 unread messages.", + "favoriteFruit": "strawberry" + }, + { + "_id": "58fa646daebd9bf06118c1cd", + "index": 118, + "guid": "d3e45591-90d0-4292-baf3-46fcc8b2cd65", + "isActive": false, + "balance": "$1,830.03", + "picture": "http://placehold.it/32x32", + "age": 23, + "eyeColor": "brown", + "name": "Navarro Rowland", + "gender": "male", + "company": "MEMORA", + "email": "navarrorowland@memora.com", + "phone": "+1 (800) 497-3371", + "address": "451 Chester Court, Wollochet, Puerto Rico, 1706", + "registered": "2016-08-12T12:21:06 -02:00", + "latitude": -66.720133, + "longitude": -157.961716, + "tags": [ + "dolore", + "deserunt", + "officia", + "do", + "irure", + "incididunt", + "mollit" + ], + "friends": [ + { + "id": 0, + "name": "Corina Barnett" + }, + { + "id": 1, + "name": "Bush Callahan" + }, + { + "id": 2, + "name": "Schroeder Hunt" + } + ], + "greeting": "Hello, Navarro Rowland! You have 5 unread messages.", + "favoriteFruit": "strawberry" + }, + { + "_id": "58fa646de87e9ecf64a2fc7c", + "index": 119, + "guid": "7c0f1a89-7509-4cd5-9ecf-508d21828c7e", + "isActive": true, + "balance": "$3,291.89", + "picture": "http://placehold.it/32x32", + "age": 21, + "eyeColor": "blue", + "name": "Aileen Carlson", + "gender": "female", + "company": "ZILODYNE", + "email": "aileencarlson@zilodyne.com", + "phone": "+1 (816) 530-3393", + "address": "645 Ridgewood Place, Yukon, Rhode Island, 6267", + "registered": "2016-05-01T10:56:26 -02:00", + "latitude": 86.715515, + "longitude": 131.479595, + "tags": [ + "aute", + "Lorem", + "reprehenderit", + "est", + "et", + "anim", + "deserunt" + ], + "friends": [ + { + "id": 0, + "name": "Johnson Wilkerson" + }, + { + "id": 1, + "name": "Henry Downs" + }, + { + "id": 2, + "name": "Sheri Reynolds" + } + ], + "greeting": "Hello, Aileen Carlson! You have 7 unread messages.", + "favoriteFruit": "apple" + }, + { + "_id": "58fa646d1f748dc5d2b1e81d", + "index": 120, + "guid": "e07b5d25-bd3e-4726-a848-608f5a5d5183", + "isActive": true, + "balance": "$2,437.96", + "picture": "http://placehold.it/32x32", + "age": 40, + "eyeColor": "brown", + "name": "Mary Adams", + "gender": "female", + "company": "RECRISYS", + "email": "maryadams@recrisys.com", + "phone": "+1 (991) 428-2518", + "address": "966 Irving Avenue, Lindisfarne, Alabama, 1218", + "registered": "2015-04-19T11:02:34 -02:00", + "latitude": -30.89009, + "longitude": -4.475556, + "tags": [ + "officia", + "incididunt", + "ullamco", + "sint", + "ipsum", + "duis", + "veniam" + ], + "friends": [ + { + "id": 0, + "name": "Goodman Castillo" + }, + { + "id": 1, + "name": "Maggie Day" + }, + { + "id": 2, + "name": "Sophie Henderson" + } + ], + "greeting": "Hello, Mary Adams! You have 3 unread messages.", + "favoriteFruit": "banana" + }, + { + "_id": "58fa646d974228abb7b1d7a7", + "index": 121, + "guid": "bfa911f9-a40e-4dab-a007-7fd4f850bda0", + "isActive": true, + "balance": "$1,587.80", + "picture": "http://placehold.it/32x32", + "age": 35, + "eyeColor": "green", + "name": "May Velasquez", + "gender": "female", + "company": "PERMADYNE", + "email": "mayvelasquez@permadyne.com", + "phone": "+1 (822) 442-3557", + "address": "812 Ira Court, Datil, Wisconsin, 9833", + "registered": "2017-04-13T07:21:37 -02:00", + "latitude": -40.988345, + "longitude": 30.507673, + "tags": [ + "labore", + "et", + "nulla", + "laboris", + "non", + "voluptate", + "laboris" + ], + "friends": [ + { + "id": 0, + "name": "Kristi Schroeder" + }, + { + "id": 1, + "name": "Augusta Michael" + }, + { + "id": 2, + "name": "Jane Douglas" + } + ], + "greeting": "Hello, May Velasquez! You have 8 unread messages.", + "favoriteFruit": "banana" + }, + { + "_id": "58fa646df82f16b178aa79c0", + "index": 122, + "guid": "9d6849aa-e260-4d31-af7c-40d7f9c767aa", + "isActive": false, + "balance": "$1,416.21", + "picture": "http://placehold.it/32x32", + "age": 30, + "eyeColor": "blue", + "name": "Gregory Whitfield", + "gender": "male", + "company": "CEDWARD", + "email": "gregorywhitfield@cedward.com", + "phone": "+1 (922) 591-2714", + "address": "887 Bush Street, Kent, California, 1938", + "registered": "2017-01-16T09:56:31 -01:00", + "latitude": 29.288822, + "longitude": -51.403759, + "tags": [ + "sit", + "Lorem", + "anim", + "aute", + "et", + "labore", + "dolore" + ], + "friends": [ + { + "id": 0, + "name": "Walton Mccormick" + }, + { + "id": 1, + "name": "Casey Mccarthy" + }, + { + "id": 2, + "name": "Rochelle Lyons" + } + ], + "greeting": "Hello, Gregory Whitfield! You have 9 unread messages.", + "favoriteFruit": "apple" + }, + { + "_id": "58fa646d49a80834bbb1cf7b", + "index": 123, + "guid": "5235f17c-022a-49b0-a2e5-1724d8044c16", + "isActive": false, + "balance": "$3,728.49", + "picture": "http://placehold.it/32x32", + "age": 35, + "eyeColor": "green", + "name": "Dena Osborne", + "gender": "female", + "company": "OPTYK", + "email": "denaosborne@optyk.com", + "phone": "+1 (836) 597-3802", + "address": "765 Schenck Court, Brady, Arkansas, 6168", + "registered": "2015-06-20T05:14:08 -02:00", + "latitude": 49.153241, + "longitude": -17.25969, + "tags": [ + "nisi", + "culpa", + "sint", + "sint", + "reprehenderit", + "sunt", + "Lorem" + ], + "friends": [ + { + "id": 0, + "name": "Jaime Frye" + }, + { + "id": 1, + "name": "Christy Briggs" + }, + { + "id": 2, + "name": "Stephenson Mccullough" + } + ], + "greeting": "Hello, Dena Osborne! You have 1 unread messages.", + "favoriteFruit": "strawberry" + }, + { + "_id": "58fa646dfbbe2e7dc05f1d63", + "index": 124, + "guid": "bb869f84-0b8d-4d99-9fe5-484f6a3e1499", + "isActive": false, + "balance": "$1,556.78", + "picture": "http://placehold.it/32x32", + "age": 33, + "eyeColor": "brown", + "name": "Karina Maldonado", + "gender": "female", + "company": "CABLAM", + "email": "karinamaldonado@cablam.com", + "phone": "+1 (854) 572-2714", + "address": "404 Fleet Street, Cobbtown, Vermont, 7560", + "registered": "2014-02-27T03:01:12 -01:00", + "latitude": -75.63638, + "longitude": 91.61689, + "tags": [ + "tempor", + "adipisicing", + "quis", + "proident", + "est", + "esse", + "commodo" + ], + "friends": [ + { + "id": 0, + "name": "Lesley Sawyer" + }, + { + "id": 1, + "name": "Esperanza Mcintosh" + }, + { + "id": 2, + "name": "Grace Lang" + } + ], + "greeting": "Hello, Karina Maldonado! You have 9 unread messages.", + "favoriteFruit": "apple" + }, + { + "_id": "58fa646d8de4e05292cdeedc", + "index": 125, + "guid": "2543c764-0c2e-4e13-a2f8-707233157de1", + "isActive": false, + "balance": "$2,293.19", + "picture": "http://placehold.it/32x32", + "age": 28, + "eyeColor": "blue", + "name": "Jerri Cherry", + "gender": "female", + "company": "FRENEX", + "email": "jerricherry@frenex.com", + "phone": "+1 (917) 461-2379", + "address": "426 Gerry Street, Crucible, New Hampshire, 1100", + "registered": "2016-05-19T10:31:59 -02:00", + "latitude": -0.999925, + "longitude": -45.241596, + "tags": [ + "nostrud", + "pariatur", + "Lorem", + "nulla", + "officia", + "eu", + "deserunt" + ], + "friends": [ + { + "id": 0, + "name": "Susana Rivas" + }, + { + "id": 1, + "name": "Melinda Olsen" + }, + { + "id": 2, + "name": "Iva Wood" + } + ], + "greeting": "Hello, Jerri Cherry! You have 7 unread messages.", + "favoriteFruit": "strawberry" + }, + { + "_id": "58fa646dbe6658b40b7cb743", + "index": 126, + "guid": "90613061-fc26-4eac-a3fd-98efd9728720", + "isActive": true, + "balance": "$3,158.19", + "picture": "http://placehold.it/32x32", + "age": 33, + "eyeColor": "blue", + "name": "Merritt Quinn", + "gender": "male", + "company": "ORBAXTER", + "email": "merrittquinn@orbaxter.com", + "phone": "+1 (831) 440-2088", + "address": "786 Cozine Avenue, Wyoming, Delaware, 3337", + "registered": "2016-05-27T03:28:50 -02:00", + "latitude": -11.404925, + "longitude": -14.733943, + "tags": [ + "nisi", + "consequat", + "qui", + "Lorem", + "ipsum", + "excepteur", + "dolor" + ], + "friends": [ + { + "id": 0, + "name": "Richardson Gross" + }, + { + "id": 1, + "name": "Foreman Mcknight" + }, + { + "id": 2, + "name": "Billie Houston" + } + ], + "greeting": "Hello, Merritt Quinn! You have 1 unread messages.", + "favoriteFruit": "banana" + }, + { + "_id": "58fa646d4f8f99526f7fe491", + "index": 127, + "guid": "e34a4536-c44f-4ff4-8aa1-4314bb0e54db", + "isActive": false, + "balance": "$3,969.54", + "picture": "http://placehold.it/32x32", + "age": 24, + "eyeColor": "green", + "name": "Murphy Adkins", + "gender": "male", + "company": "ESCENTA", + "email": "murphyadkins@escenta.com", + "phone": "+1 (888) 430-2084", + "address": "116 Clarendon Road, Martell, New Jersey, 1103", + "registered": "2016-06-08T01:22:00 -02:00", + "latitude": 28.447401, + "longitude": 145.558806, + "tags": [ + "et", + "nulla", + "Lorem", + "occaecat", + "amet", + "ea", + "laborum" + ], + "friends": [ + { + "id": 0, + "name": "Henson Terry" + }, + { + "id": 1, + "name": "Gillespie Barber" + }, + { + "id": 2, + "name": "Debora Hickman" + } + ], + "greeting": "Hello, Murphy Adkins! You have 1 unread messages.", + "favoriteFruit": "banana" + }, + { + "_id": "58fa646d9348616404edc625", + "index": 128, + "guid": "1166a6dc-8c8c-4264-a167-d33d7be6d810", + "isActive": true, + "balance": "$1,724.48", + "picture": "http://placehold.it/32x32", + "age": 37, + "eyeColor": "brown", + "name": "Marjorie Thompson", + "gender": "female", + "company": "COMVOY", + "email": "marjoriethompson@comvoy.com", + "phone": "+1 (987) 425-2315", + "address": "336 Metropolitan Avenue, Zarephath, Nebraska, 7219", + "registered": "2017-04-01T12:35:15 -02:00", + "latitude": -43.767729, + "longitude": -137.618989, + "tags": [ + "quis", + "do", + "Lorem", + "amet", + "minim", + "eu", + "reprehenderit" + ], + "friends": [ + { + "id": 0, + "name": "Skinner Bradford" + }, + { + "id": 1, + "name": "Regina Dunn" + }, + { + "id": 2, + "name": "Mandy Reed" + } + ], + "greeting": "Hello, Marjorie Thompson! You have 10 unread messages.", + "favoriteFruit": "banana" + }, + { + "_id": "58fa646d45a18ebf46e92ab0", + "index": 129, + "guid": "d23ce633-07c5-492f-85c0-e61b718ced30", + "isActive": true, + "balance": "$3,922.99", + "picture": "http://placehold.it/32x32", + "age": 24, + "eyeColor": "green", + "name": "Francis Juarez", + "gender": "male", + "company": "GLUKGLUK", + "email": "francisjuarez@glukgluk.com", + "phone": "+1 (873) 513-2456", + "address": "456 Roosevelt Court, Efland, Nevada, 2771", + "registered": "2015-12-04T03:21:31 -01:00", + "latitude": 71.856563, + "longitude": 155.687754, + "tags": [ + "laboris", + "cupidatat", + "est", + "eu", + "sit", + "deserunt", + "elit" + ], + "friends": [ + { + "id": 0, + "name": "Anthony Mcneil" + }, + { + "id": 1, + "name": "Joan Mercer" + }, + { + "id": 2, + "name": "Craig Dixon" + } + ], + "greeting": "Hello, Francis Juarez! You have 4 unread messages.", + "favoriteFruit": "apple" + }, + { + "_id": "58fa646d969073210ee7b50d", + "index": 130, + "guid": "51b71033-9d3c-4ef8-b00e-cfb28c9e1160", + "isActive": true, + "balance": "$2,612.61", + "picture": "http://placehold.it/32x32", + "age": 24, + "eyeColor": "brown", + "name": "Gloria Coleman", + "gender": "female", + "company": "EVEREST", + "email": "gloriacoleman@everest.com", + "phone": "+1 (941) 488-3290", + "address": "657 Himrod Street, Cetronia, Mississippi, 801", + "registered": "2015-11-23T04:00:31 -01:00", + "latitude": 36.061156, + "longitude": 60.009705, + "tags": [ + "velit", + "dolore", + "labore", + "incididunt", + "deserunt", + "tempor", + "nostrud" + ], + "friends": [ + { + "id": 0, + "name": "Wood Snider" + }, + { + "id": 1, + "name": "Chandra Barrera" + }, + { + "id": 2, + "name": "Laverne Skinner" + } + ], + "greeting": "Hello, Gloria Coleman! You have 9 unread messages.", + "favoriteFruit": "banana" + }, + { + "_id": "58fa646d5bcf9d5fdbfd5621", + "index": 131, + "guid": "b57eac28-6c0f-4177-9028-f83ce2e928b2", + "isActive": false, + "balance": "$3,209.11", + "picture": "http://placehold.it/32x32", + "age": 36, + "eyeColor": "brown", + "name": "Gabriela Cruz", + "gender": "female", + "company": "EARTHWAX", + "email": "gabrielacruz@earthwax.com", + "phone": "+1 (876) 539-2774", + "address": "335 Baltic Street, Gibbsville, New York, 1845", + "registered": "2017-03-01T08:56:41 -01:00", + "latitude": -82.915585, + "longitude": 82.66247, + "tags": [ + "officia", + "reprehenderit", + "eiusmod", + "minim", + "quis", + "ipsum", + "laboris" + ], + "friends": [ + { + "id": 0, + "name": "Melissa Dejesus" + }, + { + "id": 1, + "name": "Laurel Franklin" + }, + { + "id": 2, + "name": "Rios Nixon" + } + ], + "greeting": "Hello, Gabriela Cruz! You have 3 unread messages.", + "favoriteFruit": "apple" + }, + { + "_id": "58fa646da749a0a4c47bf41f", + "index": 132, + "guid": "74fc6ded-da7b-48b9-991d-2ca9f1ee1901", + "isActive": true, + "balance": "$3,730.68", + "picture": "http://placehold.it/32x32", + "age": 21, + "eyeColor": "brown", + "name": "Martinez Bartlett", + "gender": "male", + "company": "MIRACULA", + "email": "martinezbartlett@miracula.com", + "phone": "+1 (903) 400-2800", + "address": "969 Jackson Place, Vienna, Oklahoma, 3282", + "registered": "2014-09-21T12:45:43 -02:00", + "latitude": -68.778345, + "longitude": 19.094414, + "tags": [ + "proident", + "proident", + "occaecat", + "velit", + "ad", + "aute", + "do" + ], + "friends": [ + { + "id": 0, + "name": "Francisca Hatfield" + }, + { + "id": 1, + "name": "Sears Mcconnell" + }, + { + "id": 2, + "name": "Gwen Washington" + } + ], + "greeting": "Hello, Martinez Bartlett! You have 2 unread messages.", + "favoriteFruit": "banana" + }, + { + "_id": "58fa646dee42275dd32a111f", + "index": 133, + "guid": "72bef6e5-0799-478c-b486-44683b98b843", + "isActive": true, + "balance": "$2,045.99", + "picture": "http://placehold.it/32x32", + "age": 39, + "eyeColor": "blue", + "name": "Mcbride Noel", + "gender": "male", + "company": "CHILLIUM", + "email": "mcbridenoel@chillium.com", + "phone": "+1 (814) 520-2323", + "address": "149 Beaumont Street, Saddlebrooke, Maryland, 8389", + "registered": "2014-03-04T03:42:36 -01:00", + "latitude": 14.904872, + "longitude": 104.324123, + "tags": [ + "culpa", + "pariatur", + "deserunt", + "aliqua", + "fugiat", + "ipsum", + "fugiat" + ], + "friends": [ + { + "id": 0, + "name": "King Meadows" + }, + { + "id": 1, + "name": "Norton George" + }, + { + "id": 2, + "name": "Rosa Gallagher" + } + ], + "greeting": "Hello, Mcbride Noel! You have 4 unread messages.", + "favoriteFruit": "strawberry" + }, + { + "_id": "58fa646de80962391114bf06", + "index": 134, + "guid": "2a2ac27d-c158-4f5c-93b2-14d9226f5e48", + "isActive": false, + "balance": "$1,335.56", + "picture": "http://placehold.it/32x32", + "age": 29, + "eyeColor": "green", + "name": "Crawford Clay", + "gender": "male", + "company": "XIIX", + "email": "crawfordclay@xiix.com", + "phone": "+1 (844) 414-2474", + "address": "516 Portal Street, Tampico, Wyoming, 8315", + "registered": "2016-11-04T11:57:56 -01:00", + "latitude": -70.094171, + "longitude": -141.660706, + "tags": [ + "qui", + "Lorem", + "quis", + "anim", + "exercitation", + "aliqua", + "exercitation" + ], + "friends": [ + { + "id": 0, + "name": "Chen Guy" + }, + { + "id": 1, + "name": "Pickett Weaver" + }, + { + "id": 2, + "name": "Hobbs Rollins" + } + ], + "greeting": "Hello, Crawford Clay! You have 5 unread messages.", + "favoriteFruit": "strawberry" + }, + { + "_id": "58fa646d2501ad2057470ba0", + "index": 135, + "guid": "6e0c8fbe-d8d3-47ba-afec-f73f486b6724", + "isActive": false, + "balance": "$1,677.75", + "picture": "http://placehold.it/32x32", + "age": 36, + "eyeColor": "blue", + "name": "Allison Gray", + "gender": "male", + "company": "COMTRAK", + "email": "allisongray@comtrak.com", + "phone": "+1 (978) 537-3908", + "address": "107 Beaver Street, Unionville, Arizona, 6833", + "registered": "2015-01-20T04:29:44 -01:00", + "latitude": 89.734894, + "longitude": 158.818049, + "tags": [ + "in", + "mollit", + "mollit", + "nostrud", + "nulla", + "dolore", + "veniam" + ], + "friends": [ + { + "id": 0, + "name": "Tabatha Gilbert" + }, + { + "id": 1, + "name": "Becky Walton" + }, + { + "id": 2, + "name": "Noreen Henry" + } + ], + "greeting": "Hello, Allison Gray! You have 5 unread messages.", + "favoriteFruit": "banana" + }, + { + "_id": "58fa646de518b19c075b0ae0", + "index": 136, + "guid": "e4286785-435b-41bc-b9ca-ace978f4bd35", + "isActive": false, + "balance": "$1,178.48", + "picture": "http://placehold.it/32x32", + "age": 34, + "eyeColor": "brown", + "name": "Jasmine Coffey", + "gender": "female", + "company": "STREZZO", + "email": "jasminecoffey@strezzo.com", + "phone": "+1 (944) 599-3917", + "address": "943 Fountain Avenue, Kaka, Texas, 627", + "registered": "2015-07-26T06:13:26 -02:00", + "latitude": -49.645521, + "longitude": -142.963046, + "tags": [ + "nostrud", + "mollit", + "excepteur", + "consectetur", + "nostrud", + "occaecat", + "nostrud" + ], + "friends": [ + { + "id": 0, + "name": "Wilkerson Clements" + }, + { + "id": 1, + "name": "Bishop Levy" + }, + { + "id": 2, + "name": "Wilkinson Roberson" + } + ], + "greeting": "Hello, Jasmine Coffey! You have 4 unread messages.", + "favoriteFruit": "apple" + }, + { + "_id": "58fa646dd70f9613ec9c3c0c", + "index": 137, + "guid": "811ab663-8fe2-4647-81d1-40b3d3454ba5", + "isActive": true, + "balance": "$1,964.85", + "picture": "http://placehold.it/32x32", + "age": 36, + "eyeColor": "brown", + "name": "Hoffman Ashley", + "gender": "male", + "company": "GOKO", + "email": "hoffmanashley@goko.com", + "phone": "+1 (815) 579-2880", + "address": "635 Irving Street, Eden, Connecticut, 2603", + "registered": "2015-09-12T04:44:56 -02:00", + "latitude": -55.239516, + "longitude": 118.685339, + "tags": [ + "elit", + "est", + "id", + "laboris", + "et", + "elit", + "consequat" + ], + "friends": [ + { + "id": 0, + "name": "Alice Britt" + }, + { + "id": 1, + "name": "Dianne Dennis" + }, + { + "id": 2, + "name": "Megan Salinas" + } + ], + "greeting": "Hello, Hoffman Ashley! You have 9 unread messages.", + "favoriteFruit": "strawberry" + }, + { + "_id": "58fa646d3d67feb43dd8cef4", + "index": 138, + "guid": "b37b2cc0-45b2-4ba7-8206-2e2a8cd62f60", + "isActive": false, + "balance": "$1,302.08", + "picture": "http://placehold.it/32x32", + "age": 38, + "eyeColor": "blue", + "name": "England Hendrix", + "gender": "male", + "company": "EVENTAGE", + "email": "englandhendrix@eventage.com", + "phone": "+1 (897) 412-2403", + "address": "262 Nassau Avenue, Kohatk, Maine, 7778", + "registered": "2014-06-09T10:04:03 -02:00", + "latitude": -85.667235, + "longitude": 77.283374, + "tags": [ + "exercitation", + "exercitation", + "consectetur", + "fugiat", + "dolore", + "commodo", + "ipsum" + ], + "friends": [ + { + "id": 0, + "name": "Lee Thomas" + }, + { + "id": 1, + "name": "Sampson Moses" + }, + { + "id": 2, + "name": "Holland Tanner" + } + ], + "greeting": "Hello, England Hendrix! You have 8 unread messages.", + "favoriteFruit": "banana" + }, + { + "_id": "58fa646d9d90ffcc996b54bd", + "index": 139, + "guid": "8504ebf7-b4ae-4aba-bd40-026dd6637b20", + "isActive": true, + "balance": "$2,595.54", + "picture": "http://placehold.it/32x32", + "age": 23, + "eyeColor": "blue", + "name": "Fields Charles", + "gender": "male", + "company": "BOSTONIC", + "email": "fieldscharles@bostonic.com", + "phone": "+1 (974) 408-3061", + "address": "608 Arion Place, Oley, Idaho, 1058", + "registered": "2016-07-26T08:29:26 -02:00", + "latitude": -66.951316, + "longitude": 43.649635, + "tags": [ + "excepteur", + "laborum", + "excepteur", + "proident", + "proident", + "ipsum", + "sunt" + ], + "friends": [ + { + "id": 0, + "name": "Freeman Mays" + }, + { + "id": 1, + "name": "Morton Golden" + }, + { + "id": 2, + "name": "Joanne Montoya" + } + ], + "greeting": "Hello, Fields Charles! You have 10 unread messages.", + "favoriteFruit": "strawberry" + }, + { + "_id": "58fa646d99e60c2345fe6385", + "index": 140, + "guid": "006dee7a-cb27-415d-99ed-182611d1ccca", + "isActive": true, + "balance": "$2,085.41", + "picture": "http://placehold.it/32x32", + "age": 37, + "eyeColor": "brown", + "name": "Vickie Murphy", + "gender": "female", + "company": "MAGMINA", + "email": "vickiemurphy@magmina.com", + "phone": "+1 (973) 516-3520", + "address": "945 Ferry Place, Campo, Hawaii, 7114", + "registered": "2014-12-22T04:23:19 -01:00", + "latitude": 80.556246, + "longitude": 43.91025, + "tags": [ + "non", + "sunt", + "irure", + "ea", + "culpa", + "ipsum", + "consequat" + ], + "friends": [ + { + "id": 0, + "name": "Harriett Slater" + }, + { + "id": 1, + "name": "Adrian Lloyd" + }, + { + "id": 2, + "name": "Marsha Bauer" + } + ], + "greeting": "Hello, Vickie Murphy! You have 6 unread messages.", + "favoriteFruit": "banana" + }, + { + "_id": "58fa646d1ec1f539698644ac", + "index": 141, + "guid": "b7fee300-85cc-4333-b59c-b2d02cea02d6", + "isActive": true, + "balance": "$2,029.76", + "picture": "http://placehold.it/32x32", + "age": 35, + "eyeColor": "brown", + "name": "English Kim", + "gender": "male", + "company": "HOMELUX", + "email": "englishkim@homelux.com", + "phone": "+1 (825) 516-2952", + "address": "780 Karweg Place, Snyderville, Alaska, 7415", + "registered": "2017-01-01T04:21:23 -01:00", + "latitude": -23.727119, + "longitude": 6.123072, + "tags": [ + "ea", + "deserunt", + "cupidatat", + "culpa", + "adipisicing", + "duis", + "minim" + ], + "friends": [ + { + "id": 0, + "name": "Savage Huff" + }, + { + "id": 1, + "name": "Ila Mercado" + }, + { + "id": 2, + "name": "Claudette Marks" + } + ], + "greeting": "Hello, English Kim! You have 2 unread messages.", + "favoriteFruit": "apple" + }, + { + "_id": "58fa646d185f4ec15673adc8", + "index": 142, + "guid": "ffaf12f0-a1be-4233-ba26-be5bcf51aec7", + "isActive": false, + "balance": "$2,240.29", + "picture": "http://placehold.it/32x32", + "age": 32, + "eyeColor": "green", + "name": "Baldwin Cash", + "gender": "male", + "company": "TERASCAPE", + "email": "baldwincash@terascape.com", + "phone": "+1 (972) 551-3447", + "address": "926 Landis Court, Lithium, Iowa, 5219", + "registered": "2015-04-05T06:34:38 -02:00", + "latitude": -53.335997, + "longitude": 85.341267, + "tags": [ + "amet", + "pariatur", + "laborum", + "velit", + "nulla", + "cupidatat", + "reprehenderit" + ], + "friends": [ + { + "id": 0, + "name": "Cassandra Bender" + }, + { + "id": 1, + "name": "Crystal Mann" + }, + { + "id": 2, + "name": "Delaney Nolan" + } + ], + "greeting": "Hello, Baldwin Cash! You have 6 unread messages.", + "favoriteFruit": "apple" + }, + { + "_id": "58fa646ddcbda994c361b679", + "index": 143, + "guid": "face4851-4828-42e8-b287-8be9e93f1310", + "isActive": true, + "balance": "$3,989.45", + "picture": "http://placehold.it/32x32", + "age": 34, + "eyeColor": "brown", + "name": "Rosella Maddox", + "gender": "female", + "company": "MENBRAIN", + "email": "rosellamaddox@menbrain.com", + "phone": "+1 (936) 492-2122", + "address": "285 Front Street, Savage, Michigan, 3482", + "registered": "2014-04-13T01:34:13 -02:00", + "latitude": -13.764801, + "longitude": -80.684874, + "tags": [ + "quis", + "ipsum", + "ut", + "exercitation", + "mollit", + "ullamco", + "irure" + ], + "friends": [ + { + "id": 0, + "name": "Randi Weiss" + }, + { + "id": 1, + "name": "Rosanne Gould" + }, + { + "id": 2, + "name": "Sherrie York" + } + ], + "greeting": "Hello, Rosella Maddox! You have 7 unread messages.", + "favoriteFruit": "strawberry" + }, + { + "_id": "58fa646ddb703dc4048b02e3", + "index": 144, + "guid": "ccd9bf2b-dca5-4525-93a3-7602bba5f217", + "isActive": true, + "balance": "$1,427.62", + "picture": "http://placehold.it/32x32", + "age": 28, + "eyeColor": "brown", + "name": "Minnie Haney", + "gender": "female", + "company": "ZERBINA", + "email": "minniehaney@zerbina.com", + "phone": "+1 (807) 536-2469", + "address": "316 Fillmore Avenue, Orick, North Dakota, 5283", + "registered": "2016-09-26T11:59:18 -02:00", + "latitude": 72.525636, + "longitude": -163.72616, + "tags": [ + "tempor", + "consectetur", + "sunt", + "qui", + "pariatur", + "laboris", + "enim" + ], + "friends": [ + { + "id": 0, + "name": "Kendra Cote" + }, + { + "id": 1, + "name": "Barnett Mooney" + }, + { + "id": 2, + "name": "Victoria Clarke" + } + ], + "greeting": "Hello, Minnie Haney! You have 9 unread messages.", + "favoriteFruit": "banana" + }, + { + "_id": "58fa646d7fee0f9e3f5383a8", + "index": 145, + "guid": "71519050-a3db-4a64-840f-79224c632212", + "isActive": true, + "balance": "$3,978.52", + "picture": "http://placehold.it/32x32", + "age": 33, + "eyeColor": "green", + "name": "Shepard Wade", + "gender": "male", + "company": "PAPRIKUT", + "email": "shepardwade@paprikut.com", + "phone": "+1 (907) 565-2766", + "address": "840 Stuyvesant Avenue, Starks, Guam, 8253", + "registered": "2015-05-03T12:04:36 -02:00", + "latitude": -84.956066, + "longitude": 142.041046, + "tags": [ + "eiusmod", + "deserunt", + "est", + "excepteur", + "irure", + "voluptate", + "est" + ], + "friends": [ + { + "id": 0, + "name": "Kim Watson" + }, + { + "id": 1, + "name": "Adela Poole" + }, + { + "id": 2, + "name": "Bryan Klein" + } + ], + "greeting": "Hello, Shepard Wade! You have 4 unread messages.", + "favoriteFruit": "banana" + }, + { + "_id": "58fa646d426fd102caed459b", + "index": 146, + "guid": "0630180a-6838-43c3-b06b-4dbe0580494c", + "isActive": false, + "balance": "$2,518.27", + "picture": "http://placehold.it/32x32", + "age": 24, + "eyeColor": "green", + "name": "Madelyn Foley", + "gender": "female", + "company": "ERSUM", + "email": "madelynfoley@ersum.com", + "phone": "+1 (941) 536-3320", + "address": "622 Bridgewater Street, Helen, Virginia, 9246", + "registered": "2015-05-11T06:24:09 -02:00", + "latitude": -39.823293, + "longitude": 61.43805, + "tags": [ + "eiusmod", + "est", + "culpa", + "eu", + "ut", + "pariatur", + "dolore" + ], + "friends": [ + { + "id": 0, + "name": "Shelley Love" + }, + { + "id": 1, + "name": "Chaney Baker" + }, + { + "id": 2, + "name": "Jacobson Branch" + } + ], + "greeting": "Hello, Madelyn Foley! You have 1 unread messages.", + "favoriteFruit": "apple" + }, + { + "_id": "58fa646dae4346f0c65d43ea", + "index": 147, + "guid": "e5faba13-38be-4993-a520-543857f619a1", + "isActive": false, + "balance": "$2,671.45", + "picture": "http://placehold.it/32x32", + "age": 23, + "eyeColor": "brown", + "name": "Stewart Hansen", + "gender": "male", + "company": "PANZENT", + "email": "stewarthansen@panzent.com", + "phone": "+1 (866) 579-2453", + "address": "156 Highland Avenue, Cecilia, West Virginia, 421", + "registered": "2014-09-29T01:43:21 -02:00", + "latitude": -61.60382, + "longitude": -12.976763, + "tags": [ + "eu", + "voluptate", + "mollit", + "irure", + "consequat", + "sint", + "mollit" + ], + "friends": [ + { + "id": 0, + "name": "Jean Lowery" + }, + { + "id": 1, + "name": "Collier Weeks" + }, + { + "id": 2, + "name": "Hinton Keith" + } + ], + "greeting": "Hello, Stewart Hansen! You have 3 unread messages.", + "favoriteFruit": "apple" + }, + { + "_id": "58fa646d3d48dcd127ab23ae", + "index": 148, + "guid": "1d6fc438-e5b2-4a6f-bec2-cd6384ca2786", + "isActive": true, + "balance": "$1,360.21", + "picture": "http://placehold.it/32x32", + "age": 33, + "eyeColor": "brown", + "name": "Douglas Sampson", + "gender": "male", + "company": "INFOTRIPS", + "email": "douglassampson@infotrips.com", + "phone": "+1 (999) 484-3424", + "address": "465 Louisiana Avenue, Macdona, Kentucky, 8476", + "registered": "2015-11-22T03:28:38 -01:00", + "latitude": -38.252216, + "longitude": 124.562361, + "tags": [ + "aute", + "deserunt", + "enim", + "minim", + "nostrud", + "ut", + "excepteur" + ], + "friends": [ + { + "id": 0, + "name": "Charmaine Lester" + }, + { + "id": 1, + "name": "Rogers Gonzalez" + }, + { + "id": 2, + "name": "Kimberly Chapman" + } + ], + "greeting": "Hello, Douglas Sampson! You have 3 unread messages.", + "favoriteFruit": "strawberry" + }, + { + "_id": "58fa646d462247398c6aef2d", + "index": 149, + "guid": "9d084a37-7e5c-4a53-9eb0-c1943f14165a", + "isActive": true, + "balance": "$1,027.93", + "picture": "http://placehold.it/32x32", + "age": 28, + "eyeColor": "blue", + "name": "Isabel Moreno", + "gender": "female", + "company": "GOLOGY", + "email": "isabelmoreno@gology.com", + "phone": "+1 (998) 438-2963", + "address": "838 Pershing Loop, Wiscon, Northern Mariana Islands, 9743", + "registered": "2014-01-10T02:10:12 -01:00", + "latitude": -8.270891, + "longitude": -30.557404, + "tags": [ + "aliquip", + "exercitation", + "non", + "duis", + "deserunt", + "fugiat", + "duis" + ], + "friends": [ + { + "id": 0, + "name": "Robert Bond" + }, + { + "id": 1, + "name": "Nunez Rutledge" + }, + { + "id": 2, + "name": "Vicki Foreman" + } + ], + "greeting": "Hello, Isabel Moreno! You have 4 unread messages.", + "favoriteFruit": "banana" + }, + { + "_id": "58fa646d711604a84ad48037", + "index": 150, + "guid": "7f7235be-3b2e-40de-9aba-f9a4444912f9", + "isActive": true, + "balance": "$2,887.80", + "picture": "http://placehold.it/32x32", + "age": 36, + "eyeColor": "brown", + "name": "Daniel Gillespie", + "gender": "male", + "company": "ACCIDENCY", + "email": "danielgillespie@accidency.com", + "phone": "+1 (834) 492-2226", + "address": "463 Bragg Court, Snowville, Virgin Islands, 6188", + "registered": "2015-01-14T04:24:23 -01:00", + "latitude": -71.98587, + "longitude": -126.325614, + "tags": [ + "ad", + "ea", + "velit", + "nostrud", + "aute", + "excepteur", + "sint" + ], + "friends": [ + { + "id": 0, + "name": "Jeannine Ware" + }, + { + "id": 1, + "name": "Hood Sparks" + }, + { + "id": 2, + "name": "Kathryn Beck" + } + ], + "greeting": "Hello, Daniel Gillespie! You have 4 unread messages.", + "favoriteFruit": "banana" + }, + { + "_id": "58fa646de64068765df25520", + "index": 151, + "guid": "86bc8f3e-332d-4d0a-9060-769dc7e2553f", + "isActive": false, + "balance": "$2,657.55", + "picture": "http://placehold.it/32x32", + "age": 35, + "eyeColor": "green", + "name": "Rhonda Doyle", + "gender": "female", + "company": "ENERSOL", + "email": "rhondadoyle@enersol.com", + "phone": "+1 (809) 568-3005", + "address": "462 Wyckoff Avenue, Elliott, Federated States Of Micronesia, 1637", + "registered": "2015-07-23T02:00:02 -02:00", + "latitude": -35.594066, + "longitude": 11.592149, + "tags": [ + "mollit", + "mollit", + "ad", + "magna", + "enim", + "ex", + "nostrud" + ], + "friends": [ + { + "id": 0, + "name": "Diane Bentley" + }, + { + "id": 1, + "name": "Melva Kirk" + }, + { + "id": 2, + "name": "Thelma Jensen" + } + ], + "greeting": "Hello, Rhonda Doyle! You have 9 unread messages.", + "favoriteFruit": "strawberry" + }, + { + "_id": "58fa646d9301427bfd82c26d", + "index": 152, + "guid": "4c7bed04-1804-48ef-ae87-bca7143d92e6", + "isActive": false, + "balance": "$2,992.49", + "picture": "http://placehold.it/32x32", + "age": 38, + "eyeColor": "green", + "name": "Marquita Rasmussen", + "gender": "female", + "company": "ISOSPHERE", + "email": "marquitarasmussen@isosphere.com", + "phone": "+1 (806) 576-3332", + "address": "881 Vandam Street, Crown, Utah, 1683", + "registered": "2016-01-28T12:56:47 -01:00", + "latitude": -0.320386, + "longitude": -38.827551, + "tags": [ + "sint", + "qui", + "consequat", + "irure", + "laboris", + "quis", + "ut" + ], + "friends": [ + { + "id": 0, + "name": "Willis Tran" + }, + { + "id": 1, + "name": "Bobbie Pitts" + }, + { + "id": 2, + "name": "Ellis Owen" + } + ], + "greeting": "Hello, Marquita Rasmussen! You have 3 unread messages.", + "favoriteFruit": "apple" + }, + { + "_id": "58fa646de048e52be9f9b99e", + "index": 153, + "guid": "8a06615c-7d6a-4fe3-ab22-bc72dd78906f", + "isActive": false, + "balance": "$3,196.26", + "picture": "http://placehold.it/32x32", + "age": 21, + "eyeColor": "green", + "name": "Nichole Fowler", + "gender": "female", + "company": "ZYTRAC", + "email": "nicholefowler@zytrac.com", + "phone": "+1 (828) 529-2400", + "address": "760 Graham Avenue, Leeper, Georgia, 6268", + "registered": "2014-05-27T09:42:47 -02:00", + "latitude": -22.58195, + "longitude": -153.586714, + "tags": [ + "id", + "occaecat", + "duis", + "aliqua", + "officia", + "adipisicing", + "sunt" + ], + "friends": [ + { + "id": 0, + "name": "Candy Hinton" + }, + { + "id": 1, + "name": "Dale Knowles" + }, + { + "id": 2, + "name": "Cooke Best" + } + ], + "greeting": "Hello, Nichole Fowler! You have 10 unread messages.", + "favoriteFruit": "strawberry" + }, + { + "_id": "58fa646d1520fce68830fcbd", + "index": 154, + "guid": "2b4d361a-fc92-4118-ba13-532c11bee3a9", + "isActive": false, + "balance": "$3,849.97", + "picture": "http://placehold.it/32x32", + "age": 35, + "eyeColor": "blue", + "name": "Kelley Patton", + "gender": "male", + "company": "OMATOM", + "email": "kelleypatton@omatom.com", + "phone": "+1 (850) 533-3252", + "address": "672 Lincoln Road, Biddle, Pennsylvania, 4619", + "registered": "2016-11-13T11:30:45 -01:00", + "latitude": 9.500461, + "longitude": 178.187655, + "tags": [ + "nulla", + "aliquip", + "commodo", + "minim", + "fugiat", + "commodo", + "amet" + ], + "friends": [ + { + "id": 0, + "name": "Rich Schultz" + }, + { + "id": 1, + "name": "Holmes Jennings" + }, + { + "id": 2, + "name": "Bell Odom" + } + ], + "greeting": "Hello, Kelley Patton! You have 3 unread messages.", + "favoriteFruit": "banana" + }, + { + "_id": "58fa646d07c004be5a5796ad", + "index": 155, + "guid": "e24debec-75b9-4b51-88b5-88882fc29aa7", + "isActive": false, + "balance": "$3,000.00", + "picture": "http://placehold.it/32x32", + "age": 40, + "eyeColor": "brown", + "name": "Webster Webster", + "gender": "male", + "company": "EQUITOX", + "email": "websterwebster@equitox.com", + "phone": "+1 (976) 518-2859", + "address": "369 Summit Street, Frystown, Indiana, 5246", + "registered": "2014-03-28T02:48:31 -01:00", + "latitude": -40.54358, + "longitude": 2.4346, + "tags": [ + "pariatur", + "eu", + "quis", + "magna", + "laboris", + "ad", + "id" + ], + "friends": [ + { + "id": 0, + "name": "Mcintosh Browning" + }, + { + "id": 1, + "name": "Patrica Vaughan" + }, + { + "id": 2, + "name": "Frieda Bryan" + } + ], + "greeting": "Hello, Webster Webster! You have 9 unread messages.", + "favoriteFruit": "apple" + }, + { + "_id": "58fa646dddb9d2376a599c69", + "index": 156, + "guid": "063c0a90-896c-477a-bc27-8980d7065700", + "isActive": false, + "balance": "$1,459.93", + "picture": "http://placehold.it/32x32", + "age": 20, + "eyeColor": "blue", + "name": "Duran Chambers", + "gender": "male", + "company": "MANTRO", + "email": "duranchambers@mantro.com", + "phone": "+1 (925) 457-2254", + "address": "361 Powell Street, Eureka, District Of Columbia, 1819", + "registered": "2016-05-08T06:11:41 -02:00", + "latitude": -76.255407, + "longitude": 75.868121, + "tags": [ + "aliquip", + "incididunt", + "exercitation", + "enim", + "non", + "labore", + "minim" + ], + "friends": [ + { + "id": 0, + "name": "Edith Elliott" + }, + { + "id": 1, + "name": "Cunningham Grimes" + }, + { + "id": 2, + "name": "Best Patrick" + } + ], + "greeting": "Hello, Duran Chambers! You have 9 unread messages.", + "favoriteFruit": "banana" + }, + { + "_id": "58fa646dfe2d62b548d6c991", + "index": 157, + "guid": "929e90e3-a260-4ed3-8063-b6e6e17ba461", + "isActive": false, + "balance": "$3,184.74", + "picture": "http://placehold.it/32x32", + "age": 27, + "eyeColor": "green", + "name": "Jacquelyn Norman", + "gender": "female", + "company": "ZOSIS", + "email": "jacquelynnorman@zosis.com", + "phone": "+1 (802) 520-3489", + "address": "803 Martense Street, Brownlee, South Dakota, 8721", + "registered": "2014-01-09T12:49:36 -01:00", + "latitude": 30.781984, + "longitude": -46.984311, + "tags": [ + "officia", + "laborum", + "nostrud", + "laboris", + "cupidatat", + "aliquip", + "eiusmod" + ], + "friends": [ + { + "id": 0, + "name": "Durham Nielsen" + }, + { + "id": 1, + "name": "Conrad Walters" + }, + { + "id": 2, + "name": "Lauri Jackson" + } + ], + "greeting": "Hello, Jacquelyn Norman! You have 7 unread messages.", + "favoriteFruit": "apple" + }, + { + "_id": "58fa646dd4704bf11b5b96e4", + "index": 158, + "guid": "1d1c9134-6d21-4b1b-8f5b-f3871b1e1d0a", + "isActive": false, + "balance": "$3,726.88", + "picture": "http://placehold.it/32x32", + "age": 21, + "eyeColor": "brown", + "name": "Booker Bird", + "gender": "male", + "company": "TURNABOUT", + "email": "bookerbird@turnabout.com", + "phone": "+1 (845) 481-3764", + "address": "995 Guernsey Street, Volta, Ohio, 974", + "registered": "2016-01-23T03:18:30 -01:00", + "latitude": 56.197498, + "longitude": 20.6965, + "tags": [ + "tempor", + "Lorem", + "laborum", + "nostrud", + "dolore", + "in", + "aliqua" + ], + "friends": [ + { + "id": 0, + "name": "Kelley May" + }, + { + "id": 1, + "name": "Annabelle Wise" + }, + { + "id": 2, + "name": "Cabrera Carpenter" + } + ], + "greeting": "Hello, Booker Bird! You have 9 unread messages.", + "favoriteFruit": "banana" + }, + { + "_id": "58fa646d2d7906af49e4aebe", + "index": 159, + "guid": "7953ff27-6c49-432c-a2dd-adfcd930f161", + "isActive": true, + "balance": "$2,766.16", + "picture": "http://placehold.it/32x32", + "age": 26, + "eyeColor": "blue", + "name": "Macias Smith", + "gender": "male", + "company": "PLASMOX", + "email": "maciassmith@plasmox.com", + "phone": "+1 (849) 495-2165", + "address": "284 Commerce Street, Ferney, Palau, 5777", + "registered": "2016-04-19T03:29:10 -02:00", + "latitude": 51.019064, + "longitude": 79.616789, + "tags": [ + "duis", + "dolore", + "adipisicing", + "id", + "ipsum", + "consectetur", + "tempor" + ], + "friends": [ + { + "id": 0, + "name": "Kirsten Burton" + }, + { + "id": 1, + "name": "Valdez Goff" + }, + { + "id": 2, + "name": "Wilda Lindsay" + } + ], + "greeting": "Hello, Macias Smith! You have 2 unread messages.", + "favoriteFruit": "strawberry" + }, + { + "_id": "58fa646d296506d8371104b2", + "index": 160, + "guid": "bfc86dd1-5f65-42c8-9238-28b79413a969", + "isActive": true, + "balance": "$1,756.83", + "picture": "http://placehold.it/32x32", + "age": 20, + "eyeColor": "green", + "name": "Lester Barr", + "gender": "male", + "company": "DEEPENDS", + "email": "lesterbarr@deepends.com", + "phone": "+1 (896) 429-3088", + "address": "377 Fay Court, Kula, Louisiana, 3550", + "registered": "2016-12-09T08:12:50 -01:00", + "latitude": 69.993026, + "longitude": -83.356439, + "tags": [ + "nostrud", + "esse", + "enim", + "deserunt", + "culpa", + "enim", + "duis" + ], + "friends": [ + { + "id": 0, + "name": "Martha Buckner" + }, + { + "id": 1, + "name": "Burke Humphrey" + }, + { + "id": 2, + "name": "Wade Benjamin" + } + ], + "greeting": "Hello, Lester Barr! You have 10 unread messages.", + "favoriteFruit": "banana" + }, + { + "_id": "58fa646dcb0180d9140d170e", + "index": 161, + "guid": "63c3b886-9579-40f1-91f8-5fd7c581fbd3", + "isActive": true, + "balance": "$3,430.98", + "picture": "http://placehold.it/32x32", + "age": 38, + "eyeColor": "blue", + "name": "Larson Alexander", + "gender": "male", + "company": "COMDOM", + "email": "larsonalexander@comdom.com", + "phone": "+1 (908) 542-3861", + "address": "154 Ashland Place, Retsof, Montana, 6248", + "registered": "2016-05-27T06:12:19 -02:00", + "latitude": 86.872886, + "longitude": 141.466748, + "tags": [ + "dolore", + "Lorem", + "consectetur", + "aute", + "anim", + "occaecat", + "fugiat" + ], + "friends": [ + { + "id": 0, + "name": "Suarez Baldwin" + }, + { + "id": 1, + "name": "Reed Cabrera" + }, + { + "id": 2, + "name": "Russo Fry" + } + ], + "greeting": "Hello, Larson Alexander! You have 5 unread messages.", + "favoriteFruit": "strawberry" + }, + { + "_id": "58fa646d222eb3e2e6aecc71", + "index": 162, + "guid": "a1349cc8-77ac-471a-89eb-c096eb77f31d", + "isActive": true, + "balance": "$3,958.35", + "picture": "http://placehold.it/32x32", + "age": 22, + "eyeColor": "green", + "name": "Clarice White", + "gender": "female", + "company": "SNIPS", + "email": "claricewhite@snips.com", + "phone": "+1 (806) 506-3688", + "address": "523 Seaview Avenue, Bendon, Illinois, 8263", + "registered": "2016-12-24T08:41:40 -01:00", + "latitude": -89.771387, + "longitude": -61.140301, + "tags": [ + "eiusmod", + "dolor", + "veniam", + "id", + "irure", + "quis", + "id" + ], + "friends": [ + { + "id": 0, + "name": "Elinor Zimmerman" + }, + { + "id": 1, + "name": "Brittney Tillman" + }, + { + "id": 2, + "name": "Lorie Booth" + } + ], + "greeting": "Hello, Clarice White! You have 7 unread messages.", + "favoriteFruit": "strawberry" + }, + { + "_id": "58fa646d5427c2292655fd1f", + "index": 163, + "guid": "5719688e-09df-49f6-8891-ff9a6c564cfa", + "isActive": false, + "balance": "$1,815.54", + "picture": "http://placehold.it/32x32", + "age": 22, + "eyeColor": "blue", + "name": "Farrell Frederick", + "gender": "male", + "company": "ISOLOGICS", + "email": "farrellfrederick@isologics.com", + "phone": "+1 (985) 581-2045", + "address": "875 Kenmore Terrace, Sperryville, Minnesota, 431", + "registered": "2015-05-31T10:00:53 -02:00", + "latitude": -13.360299, + "longitude": 51.529498, + "tags": [ + "dolor", + "proident", + "ex", + "duis", + "ut", + "labore", + "sint" + ], + "friends": [ + { + "id": 0, + "name": "Gibbs Sweet" + }, + { + "id": 1, + "name": "Corinne Mason" + }, + { + "id": 2, + "name": "Naomi Wong" + } + ], + "greeting": "Hello, Farrell Frederick! You have 4 unread messages.", + "favoriteFruit": "strawberry" + }, + { + "_id": "58fa646d8164ce6008c29947", + "index": 164, + "guid": "aa96ef2c-dc50-469f-abd8-4425fab574d4", + "isActive": false, + "balance": "$2,646.94", + "picture": "http://placehold.it/32x32", + "age": 23, + "eyeColor": "green", + "name": "Cote Mccarty", + "gender": "male", + "company": "INSURON", + "email": "cotemccarty@insuron.com", + "phone": "+1 (862) 440-2146", + "address": "159 Lester Court, Ebro, South Carolina, 9738", + "registered": "2015-11-13T01:28:30 -01:00", + "latitude": -62.439524, + "longitude": -3.430019, + "tags": [ + "nisi", + "ut", + "eu", + "labore", + "ullamco", + "irure", + "voluptate" + ], + "friends": [ + { + "id": 0, + "name": "Audra Mcdonald" + }, + { + "id": 1, + "name": "Molly Stanley" + }, + { + "id": 2, + "name": "Gamble Abbott" + } + ], + "greeting": "Hello, Cote Mccarty! You have 10 unread messages.", + "favoriteFruit": "strawberry" + }, + { + "_id": "58fa646d6d060dae4e88b409", + "index": 165, + "guid": "0cae8afa-0cf7-4d72-89c1-7b79d5586dd9", + "isActive": false, + "balance": "$2,752.86", + "picture": "http://placehold.it/32x32", + "age": 36, + "eyeColor": "blue", + "name": "Obrien Sanchez", + "gender": "male", + "company": "PRIMORDIA", + "email": "obriensanchez@primordia.com", + "phone": "+1 (952) 586-2562", + "address": "403 Matthews Place, Delshire, New Mexico, 7603", + "registered": "2014-04-04T07:29:46 -02:00", + "latitude": 89.477255, + "longitude": 22.510272, + "tags": [ + "enim", + "aliquip", + "cupidatat", + "elit", + "ipsum", + "Lorem", + "esse" + ], + "friends": [ + { + "id": 0, + "name": "Knapp Rhodes" + }, + { + "id": 1, + "name": "Alvarez Horne" + }, + { + "id": 2, + "name": "Paula Odonnell" + } + ], + "greeting": "Hello, Obrien Sanchez! You have 9 unread messages.", + "favoriteFruit": "apple" + }, + { + "_id": "58fa646d640dc4dbc3d1427c", + "index": 166, + "guid": "645678fc-c16f-4906-ba9b-1a12a8f02228", + "isActive": false, + "balance": "$2,836.84", + "picture": "http://placehold.it/32x32", + "age": 20, + "eyeColor": "green", + "name": "Carole Molina", + "gender": "female", + "company": "TELEQUIET", + "email": "carolemolina@telequiet.com", + "phone": "+1 (943) 400-2462", + "address": "824 Independence Avenue, Clayville, North Carolina, 2628", + "registered": "2015-12-18T11:46:14 -01:00", + "latitude": 40.764993, + "longitude": 170.63064, + "tags": [ + "aute", + "sint", + "esse", + "minim", + "cupidatat", + "quis", + "est" + ], + "friends": [ + { + "id": 0, + "name": "Juana Bush" + }, + { + "id": 1, + "name": "Blevins Pate" + }, + { + "id": 2, + "name": "Ann Mcdowell" + } + ], + "greeting": "Hello, Carole Molina! You have 9 unread messages.", + "favoriteFruit": "strawberry" + }, + { + "_id": "58fa646d6bbad4afc3a8b7fa", + "index": 167, + "guid": "91ff6a9c-c856-4578-b1f8-3a0123bddbc4", + "isActive": true, + "balance": "$3,159.64", + "picture": "http://placehold.it/32x32", + "age": 26, + "eyeColor": "blue", + "name": "Maude Byers", + "gender": "female", + "company": "IMPERIUM", + "email": "maudebyers@imperium.com", + "phone": "+1 (920) 464-3619", + "address": "365 Montauk Avenue, Corriganville, Massachusetts, 6805", + "registered": "2015-02-18T02:48:30 -01:00", + "latitude": 69.345537, + "longitude": 81.940759, + "tags": [ + "pariatur", + "est", + "aliquip", + "veniam", + "velit", + "sint", + "qui" + ], + "friends": [ + { + "id": 0, + "name": "Lois Jacobs" + }, + { + "id": 1, + "name": "Frost Prince" + }, + { + "id": 2, + "name": "Johnston Obrien" + } + ], + "greeting": "Hello, Maude Byers! You have 6 unread messages.", + "favoriteFruit": "apple" + }, + { + "_id": "58fa646dc97d1bc1db00d4a3", + "index": 168, + "guid": "f14669e0-e928-4db1-8b59-0c64d3aeb4fa", + "isActive": false, + "balance": "$1,626.66", + "picture": "http://placehold.it/32x32", + "age": 23, + "eyeColor": "blue", + "name": "Elvira Swanson", + "gender": "female", + "company": "HONOTRON", + "email": "elviraswanson@honotron.com", + "phone": "+1 (858) 546-2617", + "address": "200 Jerome Street, Walton, Kansas, 4099", + "registered": "2014-01-13T03:58:22 -01:00", + "latitude": 41.153645, + "longitude": 123.92863, + "tags": [ + "culpa", + "cillum", + "duis", + "voluptate", + "elit", + "elit", + "laboris" + ], + "friends": [ + { + "id": 0, + "name": "Shelton Tyson" + }, + { + "id": 1, + "name": "Ofelia Chase" + }, + { + "id": 2, + "name": "Collins Wolfe" + } + ], + "greeting": "Hello, Elvira Swanson! You have 8 unread messages.", + "favoriteFruit": "apple" + }, + { + "_id": "58fa646dacbfa789bf04c721", + "index": 169, + "guid": "52e83db7-c1b4-4646-8ea2-e4a57aa9509c", + "isActive": false, + "balance": "$3,460.77", + "picture": "http://placehold.it/32x32", + "age": 25, + "eyeColor": "blue", + "name": "Lilian Cain", + "gender": "female", + "company": "INRT", + "email": "liliancain@inrt.com", + "phone": "+1 (927) 580-3192", + "address": "488 Driggs Avenue, Wright, American Samoa, 5322", + "registered": "2017-02-09T03:34:22 -01:00", + "latitude": 83.91725, + "longitude": -66.765869, + "tags": [ + "cupidatat", + "ea", + "velit", + "exercitation", + "culpa", + "laboris", + "commodo" + ], + "friends": [ + { + "id": 0, + "name": "Marietta Ramsey" + }, + { + "id": 1, + "name": "Kenya Lancaster" + }, + { + "id": 2, + "name": "Florine Barry" + } + ], + "greeting": "Hello, Lilian Cain! You have 9 unread messages.", + "favoriteFruit": "strawberry" + }, + { + "_id": "58fa646d6714bad847a973e8", + "index": 170, + "guid": "8e21668e-27dc-4054-8a3d-6d9e74fdad70", + "isActive": false, + "balance": "$3,257.06", + "picture": "http://placehold.it/32x32", + "age": 26, + "eyeColor": "blue", + "name": "Delores Mosley", + "gender": "female", + "company": "GEEKY", + "email": "deloresmosley@geeky.com", + "phone": "+1 (929) 544-2334", + "address": "131 Ainslie Street, Brutus, Tennessee, 867", + "registered": "2016-07-09T02:48:20 -02:00", + "latitude": 75.805369, + "longitude": 167.040311, + "tags": [ + "labore", + "ea", + "et", + "proident", + "minim", + "adipisicing", + "consectetur" + ], + "friends": [ + { + "id": 0, + "name": "Osborn Frazier" + }, + { + "id": 1, + "name": "Sargent Diaz" + }, + { + "id": 2, + "name": "Emilia Petersen" + } + ], + "greeting": "Hello, Delores Mosley! You have 5 unread messages.", + "favoriteFruit": "banana" + }, + { + "_id": "58fa646d09df4c99d25f341b", + "index": 171, + "guid": "366bea0f-97e3-4500-98d4-fdf4136d6038", + "isActive": false, + "balance": "$1,997.24", + "picture": "http://placehold.it/32x32", + "age": 38, + "eyeColor": "blue", + "name": "Hannah Jimenez", + "gender": "female", + "company": "CYCLONICA", + "email": "hannahjimenez@cyclonica.com", + "phone": "+1 (904) 447-2384", + "address": "401 Tennis Court, Holcombe, Florida, 7636", + "registered": "2014-02-08T04:26:33 -01:00", + "latitude": 79.176971, + "longitude": -118.290837, + "tags": [ + "ut", + "velit", + "commodo", + "nulla", + "velit", + "est", + "proident" + ], + "friends": [ + { + "id": 0, + "name": "Kari Pace" + }, + { + "id": 1, + "name": "Terrell Gibbs" + }, + { + "id": 2, + "name": "Cathy Rush" + } + ], + "greeting": "Hello, Hannah Jimenez! You have 10 unread messages.", + "favoriteFruit": "banana" + }, + { + "_id": "58fa646d195d9a0b6e7ac209", + "index": 172, + "guid": "b392e415-d30a-4f17-8206-99eccd969edb", + "isActive": true, + "balance": "$1,516.27", + "picture": "http://placehold.it/32x32", + "age": 40, + "eyeColor": "brown", + "name": "Howell Vincent", + "gender": "male", + "company": "HALAP", + "email": "howellvincent@halap.com", + "phone": "+1 (905) 472-2192", + "address": "963 Baughman Place, Hannasville, Missouri, 319", + "registered": "2015-12-03T01:38:14 -01:00", + "latitude": -37.007271, + "longitude": -22.790033, + "tags": [ + "qui", + "dolor", + "culpa", + "enim", + "mollit", + "aliqua", + "cupidatat" + ], + "friends": [ + { + "id": 0, + "name": "Cox Castro" + }, + { + "id": 1, + "name": "Norman Freeman" + }, + { + "id": 2, + "name": "Solomon Johns" + } + ], + "greeting": "Hello, Howell Vincent! You have 9 unread messages.", + "favoriteFruit": "strawberry" + }, + { + "_id": "58fa646dab64bcbee4c59bf8", + "index": 173, + "guid": "77d1efcc-ac92-41c6-b540-8e91002d5f46", + "isActive": false, + "balance": "$3,394.76", + "picture": "http://placehold.it/32x32", + "age": 39, + "eyeColor": "blue", + "name": "Sherman Howard", + "gender": "male", + "company": "IMANT", + "email": "shermanhoward@imant.com", + "phone": "+1 (971) 426-2820", + "address": "193 Lott Avenue, Harborton, Marshall Islands, 177", + "registered": "2016-01-03T05:36:20 -01:00", + "latitude": 18.834321, + "longitude": -62.030754, + "tags": [ + "sit", + "ad", + "tempor", + "occaecat", + "eu", + "occaecat", + "adipisicing" + ], + "friends": [ + { + "id": 0, + "name": "Goff Cortez" + }, + { + "id": 1, + "name": "Mcintyre Mccall" + }, + { + "id": 2, + "name": "Lourdes Pacheco" + } + ], + "greeting": "Hello, Sherman Howard! You have 10 unread messages.", + "favoriteFruit": "strawberry" + }, + { + "_id": "58fa646dc26c9636d80096d9", + "index": 174, + "guid": "3a843f82-010b-4dc0-b36a-ec15ac71f21f", + "isActive": true, + "balance": "$2,297.01", + "picture": "http://placehold.it/32x32", + "age": 40, + "eyeColor": "brown", + "name": "Gould Anthony", + "gender": "male", + "company": "KOZGENE", + "email": "gouldanthony@kozgene.com", + "phone": "+1 (880) 537-2055", + "address": "242 Degraw Street, Bonanza, Washington, 6594", + "registered": "2014-05-27T05:13:56 -02:00", + "latitude": 44.797578, + "longitude": 52.427593, + "tags": [ + "ex", + "irure", + "nulla", + "incididunt", + "sint", + "et", + "aute" + ], + "friends": [ + { + "id": 0, + "name": "Brooks Mitchell" + }, + { + "id": 1, + "name": "Roxie Oneil" + }, + { + "id": 2, + "name": "Alta Simon" + } + ], + "greeting": "Hello, Gould Anthony! You have 10 unread messages.", + "favoriteFruit": "strawberry" + }, + { + "_id": "58fa646d7d292f2e23ac3adc", + "index": 175, + "guid": "3fded192-cfa6-4c25-b41f-74bab15d52c5", + "isActive": false, + "balance": "$1,548.01", + "picture": "http://placehold.it/32x32", + "age": 37, + "eyeColor": "brown", + "name": "Edna Burch", + "gender": "female", + "company": "ZIORE", + "email": "ednaburch@ziore.com", + "phone": "+1 (838) 466-3532", + "address": "568 Emerald Street, Eggertsville, Colorado, 6934", + "registered": "2017-03-07T09:43:27 -01:00", + "latitude": -88.750505, + "longitude": 71.538521, + "tags": [ + "dolor", + "culpa", + "id", + "ad", + "ea", + "culpa", + "veniam" + ], + "friends": [ + { + "id": 0, + "name": "Phelps Holcomb" + }, + { + "id": 1, + "name": "Ortega Mullen" + }, + { + "id": 2, + "name": "Jo Floyd" + } + ], + "greeting": "Hello, Edna Burch! You have 9 unread messages.", + "favoriteFruit": "banana" + }, + { + "_id": "58fa646def09da89d7fc13f8", + "index": 176, + "guid": "0cdcc9bf-71af-4689-a2a0-dbb59368d730", + "isActive": false, + "balance": "$3,071.30", + "picture": "http://placehold.it/32x32", + "age": 24, + "eyeColor": "green", + "name": "Juliet Harrell", + "gender": "female", + "company": "ADORNICA", + "email": "julietharrell@adornica.com", + "phone": "+1 (909) 518-3849", + "address": "751 Dekoven Court, Ezel, Puerto Rico, 6545", + "registered": "2014-08-19T09:35:42 -02:00", + "latitude": -50.225911, + "longitude": 179.910183, + "tags": [ + "mollit", + "sunt", + "sit", + "ea", + "do", + "ea", + "non" + ], + "friends": [ + { + "id": 0, + "name": "Head Jacobson" + }, + { + "id": 1, + "name": "Hayden Bonner" + }, + { + "id": 2, + "name": "Contreras Calhoun" + } + ], + "greeting": "Hello, Juliet Harrell! You have 4 unread messages.", + "favoriteFruit": "apple" + }, + { + "_id": "58fa646d552672a4ea25be57", + "index": 177, + "guid": "0f1ee392-90c6-43ff-9d71-42c05acf50aa", + "isActive": true, + "balance": "$1,958.53", + "picture": "http://placehold.it/32x32", + "age": 32, + "eyeColor": "blue", + "name": "Fanny Mcmillan", + "gender": "female", + "company": "SECURIA", + "email": "fannymcmillan@securia.com", + "phone": "+1 (867) 482-2679", + "address": "453 Llama Court, Ellerslie, Rhode Island, 7033", + "registered": "2014-10-28T07:38:50 -01:00", + "latitude": 0.67566, + "longitude": 120.489868, + "tags": [ + "non", + "pariatur", + "elit", + "sint", + "ullamco", + "aute", + "aliquip" + ], + "friends": [ + { + "id": 0, + "name": "Davenport English" + }, + { + "id": 1, + "name": "Parsons Newton" + }, + { + "id": 2, + "name": "Maddox Long" + } + ], + "greeting": "Hello, Fanny Mcmillan! You have 3 unread messages.", + "favoriteFruit": "strawberry" + }, + { + "_id": "58fa646d84dc18e75e1e8076", + "index": 178, + "guid": "cb6e393d-3295-4051-8dfa-63334ee8bedc", + "isActive": true, + "balance": "$2,659.50", + "picture": "http://placehold.it/32x32", + "age": 35, + "eyeColor": "brown", + "name": "Hicks Kane", + "gender": "male", + "company": "VANTAGE", + "email": "hickskane@vantage.com", + "phone": "+1 (836) 537-2751", + "address": "107 Homecrest Avenue, Norfolk, Alabama, 8980", + "registered": "2016-08-30T10:15:44 -02:00", + "latitude": 31.404166, + "longitude": 129.701329, + "tags": [ + "id", + "ut", + "ipsum", + "Lorem", + "id", + "laborum", + "non" + ], + "friends": [ + { + "id": 0, + "name": "Dawson Mccoy" + }, + { + "id": 1, + "name": "Dorothy Solis" + }, + { + "id": 2, + "name": "Pierce Gates" + } + ], + "greeting": "Hello, Hicks Kane! You have 3 unread messages.", + "favoriteFruit": "strawberry" + }, + { + "_id": "58fa646d732f9067b1374c8f", + "index": 179, + "guid": "937f7d93-f763-432f-9551-c21bcc304f5b", + "isActive": false, + "balance": "$3,218.21", + "picture": "http://placehold.it/32x32", + "age": 34, + "eyeColor": "brown", + "name": "Katheryn Glenn", + "gender": "female", + "company": "NIXELT", + "email": "katherynglenn@nixelt.com", + "phone": "+1 (926) 439-2415", + "address": "706 Bergen Street, Vincent, Wisconsin, 2333", + "registered": "2016-11-23T02:24:15 -01:00", + "latitude": -0.948067, + "longitude": 91.13635, + "tags": [ + "sint", + "voluptate", + "proident", + "commodo", + "ea", + "duis", + "nulla" + ], + "friends": [ + { + "id": 0, + "name": "Hutchinson Parsons" + }, + { + "id": 1, + "name": "Carrie Cochran" + }, + { + "id": 2, + "name": "Miriam Lawson" + } + ], + "greeting": "Hello, Katheryn Glenn! You have 3 unread messages.", + "favoriteFruit": "strawberry" + }, + { + "_id": "58fa646d34faf079c18eb161", + "index": 180, + "guid": "051ee8df-44a1-4e85-bf7e-9b28ea374d42", + "isActive": true, + "balance": "$2,964.31", + "picture": "http://placehold.it/32x32", + "age": 33, + "eyeColor": "green", + "name": "Mcconnell Singleton", + "gender": "male", + "company": "GLASSTEP", + "email": "mcconnellsingleton@glasstep.com", + "phone": "+1 (903) 589-2014", + "address": "814 Lancaster Avenue, Orin, California, 690", + "registered": "2015-12-13T11:33:49 -01:00", + "latitude": -33.942052, + "longitude": -104.721531, + "tags": [ + "eiusmod", + "cupidatat", + "sit", + "ea", + "anim", + "pariatur", + "id" + ], + "friends": [ + { + "id": 0, + "name": "Witt Navarro" + }, + { + "id": 1, + "name": "Keisha Knox" + }, + { + "id": 2, + "name": "Estelle Pratt" + } + ], + "greeting": "Hello, Mcconnell Singleton! You have 9 unread messages.", + "favoriteFruit": "strawberry" + }, + { + "_id": "58fa646dec401dda0c504222", + "index": 181, + "guid": "b5de7eac-2d5f-4681-a0d1-ddfbefbb5c86", + "isActive": false, + "balance": "$2,695.30", + "picture": "http://placehold.it/32x32", + "age": 26, + "eyeColor": "blue", + "name": "Myers Cooley", + "gender": "male", + "company": "ZAYA", + "email": "myerscooley@zaya.com", + "phone": "+1 (876) 595-2553", + "address": "501 Gotham Avenue, Jeff, Arkansas, 5556", + "registered": "2014-11-30T09:25:25 -01:00", + "latitude": 50.700779, + "longitude": -60.689685, + "tags": [ + "velit", + "velit", + "reprehenderit", + "veniam", + "exercitation", + "consequat", + "nulla" + ], + "friends": [ + { + "id": 0, + "name": "Barber Parks" + }, + { + "id": 1, + "name": "Walker Wall" + }, + { + "id": 2, + "name": "Wendy Morgan" + } + ], + "greeting": "Hello, Myers Cooley! You have 8 unread messages.", + "favoriteFruit": "apple" + }, + { + "_id": "58fa646d91d8f9deebb2cd64", + "index": 182, + "guid": "0d8f9341-9637-4fa8-b207-04f817c4239a", + "isActive": false, + "balance": "$2,058.74", + "picture": "http://placehold.it/32x32", + "age": 22, + "eyeColor": "blue", + "name": "Letitia Carson", + "gender": "female", + "company": "COMBOT", + "email": "letitiacarson@combot.com", + "phone": "+1 (926) 475-2710", + "address": "381 Bartlett Street, Conway, Vermont, 2660", + "registered": "2014-08-18T12:46:41 -02:00", + "latitude": -11.961839, + "longitude": 39.43472, + "tags": [ + "est", + "pariatur", + "aute", + "quis", + "et", + "adipisicing", + "dolor" + ], + "friends": [ + { + "id": 0, + "name": "Rowena Vargas" + }, + { + "id": 1, + "name": "Phyllis Serrano" + }, + { + "id": 2, + "name": "Bonner Burgess" + } + ], + "greeting": "Hello, Letitia Carson! You have 10 unread messages.", + "favoriteFruit": "apple" + }, + { + "_id": "58fa646d7d1d5c2f6f824319", + "index": 183, + "guid": "52f3b0f3-fd65-4b4d-991b-9c004c897842", + "isActive": false, + "balance": "$2,296.14", + "picture": "http://placehold.it/32x32", + "age": 38, + "eyeColor": "blue", + "name": "Mollie Hopkins", + "gender": "female", + "company": "GEOLOGIX", + "email": "molliehopkins@geologix.com", + "phone": "+1 (907) 412-3989", + "address": "594 Richards Street, Brule, New Hampshire, 6206", + "registered": "2015-03-30T11:53:45 -02:00", + "latitude": 81.263548, + "longitude": 113.753547, + "tags": [ + "incididunt", + "officia", + "sunt", + "ipsum", + "eiusmod", + "enim", + "occaecat" + ], + "friends": [ + { + "id": 0, + "name": "Samantha Estrada" + }, + { + "id": 1, + "name": "Buckley Head" + }, + { + "id": 2, + "name": "Jackson Beard" + } + ], + "greeting": "Hello, Mollie Hopkins! You have 8 unread messages.", + "favoriteFruit": "banana" + }, + { + "_id": "58fa646d31a07b3b2a6c2c86", + "index": 184, + "guid": "35c976da-760b-4837-9a80-bd554c9bee06", + "isActive": true, + "balance": "$2,183.08", + "picture": "http://placehold.it/32x32", + "age": 21, + "eyeColor": "blue", + "name": "Shelby Estes", + "gender": "female", + "company": "SENMAO", + "email": "shelbyestes@senmao.com", + "phone": "+1 (879) 579-2654", + "address": "895 Dahill Road, Bethany, Delaware, 2652", + "registered": "2015-08-29T05:08:19 -02:00", + "latitude": -34.968691, + "longitude": 43.186378, + "tags": [ + "occaecat", + "consectetur", + "elit", + "amet", + "id", + "id", + "cillum" + ], + "friends": [ + { + "id": 0, + "name": "Katina Berry" + }, + { + "id": 1, + "name": "Lottie Strickland" + }, + { + "id": 2, + "name": "Tanner Davidson" + } + ], + "greeting": "Hello, Shelby Estes! You have 1 unread messages.", + "favoriteFruit": "banana" + }, + { + "_id": "58fa646dd1fbcbbf00b72064", + "index": 185, + "guid": "655efb9a-df15-4c21-b261-40248a698978", + "isActive": false, + "balance": "$1,023.43", + "picture": "http://placehold.it/32x32", + "age": 33, + "eyeColor": "blue", + "name": "Dana Gallegos", + "gender": "female", + "company": "PHARMEX", + "email": "danagallegos@pharmex.com", + "phone": "+1 (983) 467-2723", + "address": "345 Hill Street, Sheatown, New Jersey, 7099", + "registered": "2015-05-02T06:40:49 -02:00", + "latitude": -46.669675, + "longitude": 150.540031, + "tags": [ + "deserunt", + "officia", + "duis", + "Lorem", + "ad", + "mollit", + "sunt" + ], + "friends": [ + { + "id": 0, + "name": "Schwartz Webb" + }, + { + "id": 1, + "name": "Rivas Ayers" + }, + { + "id": 2, + "name": "Millie Duke" + } + ], + "greeting": "Hello, Dana Gallegos! You have 7 unread messages.", + "favoriteFruit": "banana" + }, + { + "_id": "58fa646d940088bdb9e923ee", + "index": 186, + "guid": "cf2ecd14-c20a-41c9-92f8-1ef41e40bf02", + "isActive": false, + "balance": "$2,458.72", + "picture": "http://placehold.it/32x32", + "age": 25, + "eyeColor": "green", + "name": "Pearlie Bray", + "gender": "female", + "company": "QUIZMO", + "email": "pearliebray@quizmo.com", + "phone": "+1 (914) 592-2326", + "address": "182 Brooklyn Road, Lewis, Nebraska, 6436", + "registered": "2014-04-06T02:21:31 -02:00", + "latitude": -73.365722, + "longitude": 73.749104, + "tags": [ + "ut", + "exercitation", + "tempor", + "nostrud", + "dolor", + "officia", + "ea" + ], + "friends": [ + { + "id": 0, + "name": "Hallie Taylor" + }, + { + "id": 1, + "name": "Felecia Franks" + }, + { + "id": 2, + "name": "Jan Nicholson" + } + ], + "greeting": "Hello, Pearlie Bray! You have 4 unread messages.", + "favoriteFruit": "banana" + }, + { + "_id": "58fa646db7806c7b8c4a03cd", + "index": 187, + "guid": "294c92db-1ab0-4735-b7fc-2d230409bfe3", + "isActive": true, + "balance": "$3,000.94", + "picture": "http://placehold.it/32x32", + "age": 24, + "eyeColor": "brown", + "name": "Anastasia Byrd", + "gender": "female", + "company": "TELEPARK", + "email": "anastasiabyrd@telepark.com", + "phone": "+1 (971) 436-3790", + "address": "990 Orange Street, Gouglersville, Nevada, 5171", + "registered": "2016-05-20T05:37:37 -02:00", + "latitude": -79.124843, + "longitude": -139.318155, + "tags": [ + "cupidatat", + "Lorem", + "commodo", + "consequat", + "esse", + "minim", + "sint" + ], + "friends": [ + { + "id": 0, + "name": "Lily Donovan" + }, + { + "id": 1, + "name": "Hunter Wilder" + }, + { + "id": 2, + "name": "Winnie Cleveland" + } + ], + "greeting": "Hello, Anastasia Byrd! You have 10 unread messages.", + "favoriteFruit": "banana" + }, + { + "_id": "58fa646d1531d81083d4c0fb", + "index": 188, + "guid": "f18716d4-38cd-41e7-a9b4-7835790d7c74", + "isActive": false, + "balance": "$1,481.99", + "picture": "http://placehold.it/32x32", + "age": 26, + "eyeColor": "blue", + "name": "Dona Gibson", + "gender": "female", + "company": "ESSENSIA", + "email": "donagibson@essensia.com", + "phone": "+1 (801) 460-3086", + "address": "379 Heath Place, Lowgap, Mississippi, 1204", + "registered": "2014-05-29T02:25:26 -02:00", + "latitude": -13.240939, + "longitude": -43.843129, + "tags": [ + "eu", + "laboris", + "elit", + "nostrud", + "aliquip", + "voluptate", + "voluptate" + ], + "friends": [ + { + "id": 0, + "name": "Theresa Bradshaw" + }, + { + "id": 1, + "name": "Hatfield Vang" + }, + { + "id": 2, + "name": "Britt Rowe" + } + ], + "greeting": "Hello, Dona Gibson! You have 9 unread messages.", + "favoriteFruit": "banana" + }, + { + "_id": "58fa646d09bcf6884631318a", + "index": 189, + "guid": "3a58bb77-3a2b-49d0-bb30-abdffcc8fdef", + "isActive": true, + "balance": "$3,684.14", + "picture": "http://placehold.it/32x32", + "age": 33, + "eyeColor": "brown", + "name": "Adrienne Fuentes", + "gender": "female", + "company": "AMTAS", + "email": "adriennefuentes@amtas.com", + "phone": "+1 (955) 556-2410", + "address": "374 Roosevelt Place, Wescosville, New York, 1023", + "registered": "2017-01-30T05:32:05 -01:00", + "latitude": -11.86877, + "longitude": -155.45048, + "tags": [ + "amet", + "sint", + "eu", + "in", + "nisi", + "eiusmod", + "labore" + ], + "friends": [ + { + "id": 0, + "name": "Nanette Burks" + }, + { + "id": 1, + "name": "Patti Tate" + }, + { + "id": 2, + "name": "Beulah Mcfarland" + } + ], + "greeting": "Hello, Adrienne Fuentes! You have 6 unread messages.", + "favoriteFruit": "apple" + }, + { + "_id": "58fa646da6f9ca1bb1e2adb2", + "index": 190, + "guid": "efd0e5fb-1766-4edc-9463-5cf391683fad", + "isActive": true, + "balance": "$1,198.10", + "picture": "http://placehold.it/32x32", + "age": 22, + "eyeColor": "brown", + "name": "Arlene Kennedy", + "gender": "female", + "company": "MICROLUXE", + "email": "arlenekennedy@microluxe.com", + "phone": "+1 (857) 474-3031", + "address": "745 Rewe Street, Wawona, Oklahoma, 6566", + "registered": "2015-08-05T06:30:13 -02:00", + "latitude": 72.303966, + "longitude": 79.22199, + "tags": [ + "eiusmod", + "cillum", + "commodo", + "reprehenderit", + "consequat", + "nisi", + "ad" + ], + "friends": [ + { + "id": 0, + "name": "Barr Alvarez" + }, + { + "id": 1, + "name": "Blanca Tyler" + }, + { + "id": 2, + "name": "Gonzalez Roberts" + } + ], + "greeting": "Hello, Arlene Kennedy! You have 4 unread messages.", + "favoriteFruit": "apple" + }, + { + "_id": "58fa646d85e3374ad46645d9", + "index": 191, + "guid": "189bbf8a-ff3f-4e5b-91d4-1c99c0cf6822", + "isActive": false, + "balance": "$2,980.24", + "picture": "http://placehold.it/32x32", + "age": 28, + "eyeColor": "blue", + "name": "Figueroa Ramirez", + "gender": "male", + "company": "POLARIUM", + "email": "figueroaramirez@polarium.com", + "phone": "+1 (946) 434-3717", + "address": "916 Dakota Place, Hilltop, Maryland, 1801", + "registered": "2015-12-14T03:56:14 -01:00", + "latitude": 69.91383, + "longitude": -172.5274, + "tags": [ + "Lorem", + "deserunt", + "minim", + "enim", + "laboris", + "sunt", + "laborum" + ], + "friends": [ + { + "id": 0, + "name": "Ewing Kemp" + }, + { + "id": 1, + "name": "Callie Romero" + }, + { + "id": 2, + "name": "Glenda Sargent" + } + ], + "greeting": "Hello, Figueroa Ramirez! You have 9 unread messages.", + "favoriteFruit": "apple" + }, + { + "_id": "58fa646d6afdb265d010ee00", + "index": 192, + "guid": "314e57f8-e953-4955-ad03-1df39e6da4a6", + "isActive": true, + "balance": "$1,748.83", + "picture": "http://placehold.it/32x32", + "age": 24, + "eyeColor": "blue", + "name": "Leach Pruitt", + "gender": "male", + "company": "KOOGLE", + "email": "leachpruitt@koogle.com", + "phone": "+1 (933) 406-2038", + "address": "176 Ditmas Avenue, Sharon, Wyoming, 7683", + "registered": "2015-08-09T12:16:51 -02:00", + "latitude": 16.128254, + "longitude": -36.143072, + "tags": [ + "nulla", + "et", + "consequat", + "commodo", + "nulla", + "dolore", + "adipisicing" + ], + "friends": [ + { + "id": 0, + "name": "Luna Cervantes" + }, + { + "id": 1, + "name": "Potts Moody" + }, + { + "id": 2, + "name": "Noel Lynch" + } + ], + "greeting": "Hello, Leach Pruitt! You have 8 unread messages.", + "favoriteFruit": "strawberry" + }, + { + "_id": "58fa646d6449cbd6b8902eb0", + "index": 193, + "guid": "6c12d396-16e9-4d25-b538-4b83e8ccff9d", + "isActive": true, + "balance": "$2,697.02", + "picture": "http://placehold.it/32x32", + "age": 26, + "eyeColor": "brown", + "name": "Reyna Dudley", + "gender": "female", + "company": "EXIAND", + "email": "reynadudley@exiand.com", + "phone": "+1 (930) 488-2336", + "address": "900 Waldorf Court, Magnolia, Arizona, 1210", + "registered": "2015-07-06T04:11:22 -02:00", + "latitude": -62.794501, + "longitude": -43.61756, + "tags": [ + "ut", + "Lorem", + "ad", + "dolore", + "ea", + "labore", + "velit" + ], + "friends": [ + { + "id": 0, + "name": "Essie Melton" + }, + { + "id": 1, + "name": "Bethany Russo" + }, + { + "id": 2, + "name": "Melanie Hunter" + } + ], + "greeting": "Hello, Reyna Dudley! You have 6 unread messages.", + "favoriteFruit": "strawberry" + }, + { + "_id": "58fa646d541838dce7909106", + "index": 194, + "guid": "59bb2112-4ec5-436d-bed3-714f0f9cea31", + "isActive": false, + "balance": "$3,668.34", + "picture": "http://placehold.it/32x32", + "age": 33, + "eyeColor": "brown", + "name": "Whitaker Medina", + "gender": "male", + "company": "COFINE", + "email": "whitakermedina@cofine.com", + "phone": "+1 (941) 569-2574", + "address": "639 Randolph Street, Kylertown, Texas, 4488", + "registered": "2016-02-22T07:14:24 -01:00", + "latitude": -53.58513, + "longitude": -161.66693, + "tags": [ + "anim", + "et", + "aliquip", + "excepteur", + "id", + "ea", + "nisi" + ], + "friends": [ + { + "id": 0, + "name": "Marissa Myers" + }, + { + "id": 1, + "name": "Mckenzie Anderson" + }, + { + "id": 2, + "name": "Lorrie Mcbride" + } + ], + "greeting": "Hello, Whitaker Medina! You have 6 unread messages.", + "favoriteFruit": "strawberry" + }, + { + "_id": "58fa646df21d42091a34975d", + "index": 195, + "guid": "883d330f-3b5d-4e6b-a039-31156bc7f91d", + "isActive": true, + "balance": "$2,957.53", + "picture": "http://placehold.it/32x32", + "age": 21, + "eyeColor": "blue", + "name": "Young Harvey", + "gender": "male", + "company": "ZAGGLES", + "email": "youngharvey@zaggles.com", + "phone": "+1 (961) 591-2147", + "address": "957 Prospect Avenue, Baker, Connecticut, 5611", + "registered": "2014-08-22T10:20:40 -02:00", + "latitude": -57.519891, + "longitude": -168.868519, + "tags": [ + "excepteur", + "anim", + "aute", + "elit", + "labore", + "sunt", + "ipsum" + ], + "friends": [ + { + "id": 0, + "name": "Fran Mcfadden" + }, + { + "id": 1, + "name": "Haney Pickett" + }, + { + "id": 2, + "name": "Isabella Campbell" + } + ], + "greeting": "Hello, Young Harvey! You have 1 unread messages.", + "favoriteFruit": "strawberry" + }, + { + "_id": "58fa646d6553dc8629a339b9", + "index": 196, + "guid": "6326dfc5-68f4-443d-83fd-b342c7ce6805", + "isActive": false, + "balance": "$2,543.78", + "picture": "http://placehold.it/32x32", + "age": 35, + "eyeColor": "green", + "name": "Ratliff Spencer", + "gender": "male", + "company": "APEXIA", + "email": "ratliffspencer@apexia.com", + "phone": "+1 (892) 419-3002", + "address": "126 Bay Avenue, Lydia, Maine, 5626", + "registered": "2016-05-23T09:27:20 -02:00", + "latitude": 27.702198, + "longitude": 162.98625, + "tags": [ + "aute", + "sint", + "nisi", + "ipsum", + "minim", + "et", + "labore" + ], + "friends": [ + { + "id": 0, + "name": "Gretchen Craig" + }, + { + "id": 1, + "name": "Rene Livingston" + }, + { + "id": 2, + "name": "Candice Hoover" + } + ], + "greeting": "Hello, Ratliff Spencer! You have 1 unread messages.", + "favoriteFruit": "banana" + }, + { + "_id": "58fa646d3224af0a1ec72032", + "index": 197, + "guid": "411ac9bc-d2df-47c2-95d2-f7f7852072bc", + "isActive": true, + "balance": "$1,750.87", + "picture": "http://placehold.it/32x32", + "age": 32, + "eyeColor": "green", + "name": "Mallory Vasquez", + "gender": "female", + "company": "ZANILLA", + "email": "malloryvasquez@zanilla.com", + "phone": "+1 (958) 519-2156", + "address": "813 Autumn Avenue, Galesville, Idaho, 6524", + "registered": "2015-09-20T08:05:29 -02:00", + "latitude": 28.899004, + "longitude": 172.784211, + "tags": [ + "do", + "exercitation", + "pariatur", + "consectetur", + "dolore", + "culpa", + "nostrud" + ], + "friends": [ + { + "id": 0, + "name": "Keri Harrison" + }, + { + "id": 1, + "name": "Georgia Clemons" + }, + { + "id": 2, + "name": "Anna Evans" + } + ], + "greeting": "Hello, Mallory Vasquez! You have 10 unread messages.", + "favoriteFruit": "strawberry" + }, + { + "_id": "58fa646d4c3526f9b307ef7d", + "index": 198, + "guid": "2a95f4c4-89f6-4c77-a1c1-ae4714028211", + "isActive": false, + "balance": "$3,096.20", + "picture": "http://placehold.it/32x32", + "age": 37, + "eyeColor": "brown", + "name": "Rosa Herring", + "gender": "female", + "company": "BLEENDOT", + "email": "rosaherring@bleendot.com", + "phone": "+1 (805) 485-3691", + "address": "331 Kermit Place, Cawood, Hawaii, 4584", + "registered": "2014-09-01T05:10:51 -02:00", + "latitude": 39.318031, + "longitude": -35.412556, + "tags": [ + "consectetur", + "est", + "ut", + "nisi", + "dolore", + "exercitation", + "exercitation" + ], + "friends": [ + { + "id": 0, + "name": "Annmarie Burnett" + }, + { + "id": 1, + "name": "Pollard Snyder" + }, + { + "id": 2, + "name": "Leann Daniel" + } + ], + "greeting": "Hello, Rosa Herring! You have 9 unread messages.", + "favoriteFruit": "strawberry" + }, + { + "_id": "58fa646d173919f4404fb818", + "index": 199, + "guid": "00ad0007-016e-4a4e-b0de-03eb56c67f99", + "isActive": false, + "balance": "$2,037.60", + "picture": "http://placehold.it/32x32", + "age": 33, + "eyeColor": "green", + "name": "Luann Terrell", + "gender": "female", + "company": "PRINTSPAN", + "email": "luannterrell@printspan.com", + "phone": "+1 (981) 519-2933", + "address": "673 Hunts Lane, Southmont, Alaska, 7910", + "registered": "2014-10-11T02:58:38 -02:00", + "latitude": -43.644365, + "longitude": -116.110199, + "tags": [ + "do", + "ad", + "fugiat", + "fugiat", + "aliquip", + "deserunt", + "aute" + ], + "friends": [ + { + "id": 0, + "name": "Bridgette Blanchard" + }, + { + "id": 1, + "name": "Wyatt Leon" + }, + { + "id": 2, + "name": "Margaret Gordon" + } + ], + "greeting": "Hello, Luann Terrell! You have 8 unread messages.", + "favoriteFruit": "apple" + }, + { + "_id": "58fa646d600452c389e0c06f", + "index": 200, + "guid": "9f738e75-095f-4bc4-a4d8-27452bcd49df", + "isActive": true, + "balance": "$2,978.44", + "picture": "http://placehold.it/32x32", + "age": 28, + "eyeColor": "brown", + "name": "Clemons Henson", + "gender": "male", + "company": "SONGLINES", + "email": "clemonshenson@songlines.com", + "phone": "+1 (967) 491-3380", + "address": "754 Tillary Street, Needmore, Iowa, 4899", + "registered": "2015-07-20T07:24:15 -02:00", + "latitude": -83.572294, + "longitude": -145.946094, + "tags": [ + "minim", + "exercitation", + "voluptate", + "culpa", + "nulla", + "tempor", + "proident" + ], + "friends": [ + { + "id": 0, + "name": "Michael Allison" + }, + { + "id": 1, + "name": "Eunice Sloan" + }, + { + "id": 2, + "name": "Sweet Park" + } + ], + "greeting": "Hello, Clemons Henson! You have 6 unread messages.", + "favoriteFruit": "banana" + }, + { + "_id": "58fa646d775eed3c59997ac7", + "index": 201, + "guid": "494b26e3-5a30-4792-b9b5-0bd876f15278", + "isActive": true, + "balance": "$1,437.75", + "picture": "http://placehold.it/32x32", + "age": 27, + "eyeColor": "blue", + "name": "Janette Sykes", + "gender": "female", + "company": "HANDSHAKE", + "email": "janettesykes@handshake.com", + "phone": "+1 (880) 571-3804", + "address": "621 Withers Street, Kirk, Michigan, 616", + "registered": "2015-11-09T02:52:16 -01:00", + "latitude": 61.086887, + "longitude": 88.922545, + "tags": [ + "anim", + "ullamco", + "mollit", + "tempor", + "tempor", + "duis", + "ea" + ], + "friends": [ + { + "id": 0, + "name": "Joyce Rich" + }, + { + "id": 1, + "name": "Chase Christensen" + }, + { + "id": 2, + "name": "Rodriquez Allen" + } + ], + "greeting": "Hello, Janette Sykes! You have 1 unread messages.", + "favoriteFruit": "strawberry" + }, + { + "_id": "58fa646db88a505d5b3f83d4", + "index": 202, + "guid": "978ba0f1-c72a-4499-aab2-1af117d62249", + "isActive": true, + "balance": "$2,042.33", + "picture": "http://placehold.it/32x32", + "age": 20, + "eyeColor": "blue", + "name": "Jewell Heath", + "gender": "female", + "company": "EXOBLUE", + "email": "jewellheath@exoblue.com", + "phone": "+1 (897) 410-3388", + "address": "878 Classon Avenue, Odessa, North Dakota, 205", + "registered": "2016-02-21T11:37:06 -01:00", + "latitude": -8.124416, + "longitude": 1.222154, + "tags": [ + "ut", + "minim", + "sit", + "nisi", + "magna", + "in", + "enim" + ], + "friends": [ + { + "id": 0, + "name": "Heath Foster" + }, + { + "id": 1, + "name": "Tina Dickerson" + }, + { + "id": 2, + "name": "Burris Hawkins" + } + ], + "greeting": "Hello, Jewell Heath! You have 9 unread messages.", + "favoriteFruit": "apple" + }, + { + "_id": "58fa646d4a7331bfc6e9a9cb", + "index": 203, + "guid": "e2cfb238-2a50-431e-b75d-363b2f56eb5d", + "isActive": false, + "balance": "$2,064.95", + "picture": "http://placehold.it/32x32", + "age": 34, + "eyeColor": "brown", + "name": "Holly Battle", + "gender": "female", + "company": "REMOLD", + "email": "hollybattle@remold.com", + "phone": "+1 (853) 579-2878", + "address": "210 Moore Place, Hartsville/Hartley, Guam, 4058", + "registered": "2017-03-16T01:41:59 -01:00", + "latitude": 62.653992, + "longitude": -138.948075, + "tags": [ + "est", + "cupidatat", + "incididunt", + "ad", + "reprehenderit", + "exercitation", + "voluptate" + ], + "friends": [ + { + "id": 0, + "name": "Molina Wallace" + }, + { + "id": 1, + "name": "Lindsay Holden" + }, + { + "id": 2, + "name": "Sutton Jefferson" + } + ], + "greeting": "Hello, Holly Battle! You have 5 unread messages.", + "favoriteFruit": "strawberry" + }, + { + "_id": "58fa646db99a199e3c02d51f", + "index": 204, + "guid": "a14403af-3f41-4463-9f8d-ad1b7fc5f8c8", + "isActive": false, + "balance": "$1,061.88", + "picture": "http://placehold.it/32x32", + "age": 22, + "eyeColor": "blue", + "name": "Mitzi Hodges", + "gender": "female", + "company": "CINESANCT", + "email": "mitzihodges@cinesanct.com", + "phone": "+1 (997) 475-3286", + "address": "552 Schenectady Avenue, Coldiron, Virginia, 6308", + "registered": "2016-10-18T09:03:17 -02:00", + "latitude": -82.232983, + "longitude": -27.946093, + "tags": [ + "amet", + "proident", + "adipisicing", + "ea", + "aliqua", + "quis", + "anim" + ], + "friends": [ + { + "id": 0, + "name": "Schneider Murray" + }, + { + "id": 1, + "name": "Adele Price" + }, + { + "id": 2, + "name": "Priscilla Stephenson" + } + ], + "greeting": "Hello, Mitzi Hodges! You have 9 unread messages.", + "favoriteFruit": "apple" + }, + { + "_id": "58fa646d20c5f2216b903b79", + "index": 205, + "guid": "c113203e-ce55-40a6-9926-c9439ac51faf", + "isActive": true, + "balance": "$1,015.41", + "picture": "http://placehold.it/32x32", + "age": 26, + "eyeColor": "blue", + "name": "Ethel Strong", + "gender": "female", + "company": "PATHWAYS", + "email": "ethelstrong@pathways.com", + "phone": "+1 (909) 470-2013", + "address": "835 King Street, Berwind, West Virginia, 9960", + "registered": "2014-02-26T04:45:05 -01:00", + "latitude": 32.517892, + "longitude": -44.637377, + "tags": [ + "tempor", + "veniam", + "exercitation", + "dolor", + "mollit", + "in", + "irure" + ], + "friends": [ + { + "id": 0, + "name": "Mcdonald Forbes" + }, + { + "id": 1, + "name": "Mabel Pierce" + }, + { + "id": 2, + "name": "Celeste Little" + } + ], + "greeting": "Hello, Ethel Strong! You have 5 unread messages.", + "favoriteFruit": "strawberry" + }, + { + "_id": "58fa646db7503c4e9e5db12b", + "index": 206, + "guid": "0b39e819-ad8d-4e77-a699-d31709df3a58", + "isActive": false, + "balance": "$1,692.07", + "picture": "http://placehold.it/32x32", + "age": 36, + "eyeColor": "brown", + "name": "Cheryl Simpson", + "gender": "female", + "company": "LUNCHPOD", + "email": "cherylsimpson@lunchpod.com", + "phone": "+1 (856) 455-3937", + "address": "393 Ash Street, Logan, Kentucky, 7934", + "registered": "2016-10-29T09:55:44 -02:00", + "latitude": 81.271842, + "longitude": 107.178721, + "tags": [ + "sunt", + "proident", + "magna", + "dolor", + "magna", + "est", + "ad" + ], + "friends": [ + { + "id": 0, + "name": "Glass Velazquez" + }, + { + "id": 1, + "name": "Powell Shaw" + }, + { + "id": 2, + "name": "Madeline Nelson" + } + ], + "greeting": "Hello, Cheryl Simpson! You have 7 unread messages.", + "favoriteFruit": "banana" + }, + { + "_id": "58fa646dce282fd43760514a", + "index": 207, + "guid": "147cee7e-d993-40b4-8b4a-2ff474b2fd6c", + "isActive": false, + "balance": "$2,877.51", + "picture": "http://placehold.it/32x32", + "age": 20, + "eyeColor": "green", + "name": "Gail Avery", + "gender": "female", + "company": "TERRAGEN", + "email": "gailavery@terragen.com", + "phone": "+1 (985) 453-3835", + "address": "268 Creamer Street, Caroline, Northern Mariana Islands, 3272", + "registered": "2015-01-19T04:02:06 -01:00", + "latitude": 48.31896, + "longitude": -121.67162, + "tags": [ + "nulla", + "veniam", + "aute", + "elit", + "commodo", + "qui", + "irure" + ], + "friends": [ + { + "id": 0, + "name": "Bowman Copeland" + }, + { + "id": 1, + "name": "Perkins Greene" + }, + { + "id": 2, + "name": "Lessie Norris" + } + ], + "greeting": "Hello, Gail Avery! You have 2 unread messages.", + "favoriteFruit": "apple" + }, + { + "_id": "58fa646dac94dd1db4ef0575", + "index": 208, + "guid": "609a88ab-0df9-4630-b5f7-bb25df2eb6ae", + "isActive": true, + "balance": "$3,551.34", + "picture": "http://placehold.it/32x32", + "age": 35, + "eyeColor": "green", + "name": "Bianca Lindsey", + "gender": "female", + "company": "OVATION", + "email": "biancalindsey@ovation.com", + "phone": "+1 (831) 402-2763", + "address": "510 Tapscott Street, Brecon, Virgin Islands, 2890", + "registered": "2016-06-08T03:41:19 -02:00", + "latitude": 47.698221, + "longitude": 136.951709, + "tags": [ + "officia", + "voluptate", + "id", + "sunt", + "ut", + "ex", + "commodo" + ], + "friends": [ + { + "id": 0, + "name": "Karla Yates" + }, + { + "id": 1, + "name": "Ruiz Brock" + }, + { + "id": 2, + "name": "Swanson Bullock" + } + ], + "greeting": "Hello, Bianca Lindsey! You have 9 unread messages.", + "favoriteFruit": "banana" + }, + { + "_id": "58fa646d273db58dbd9bcd9a", + "index": 209, + "guid": "f8ab9e5a-6c7d-44cf-bedd-aa4964a6cf47", + "isActive": false, + "balance": "$1,972.81", + "picture": "http://placehold.it/32x32", + "age": 21, + "eyeColor": "blue", + "name": "Morris Hubbard", + "gender": "male", + "company": "ISBOL", + "email": "morrishubbard@isbol.com", + "phone": "+1 (984) 419-2821", + "address": "820 Clermont Avenue, Coloma, Federated States Of Micronesia, 6295", + "registered": "2015-07-06T11:11:42 -02:00", + "latitude": -89.274841, + "longitude": 137.243816, + "tags": [ + "dolor", + "ad", + "incididunt", + "aute", + "tempor", + "quis", + "duis" + ], + "friends": [ + { + "id": 0, + "name": "Duke Woodard" + }, + { + "id": 1, + "name": "Erma Raymond" + }, + { + "id": 2, + "name": "Campbell Randall" + } + ], + "greeting": "Hello, Morris Hubbard! You have 10 unread messages.", + "favoriteFruit": "strawberry" + }, + { + "_id": "58fa646dce080d33a2ef6486", + "index": 210, + "guid": "d372c4f7-9466-459b-b400-2e54e69eb6c2", + "isActive": false, + "balance": "$2,120.36", + "picture": "http://placehold.it/32x32", + "age": 24, + "eyeColor": "blue", + "name": "Prince Padilla", + "gender": "male", + "company": "GEOFARM", + "email": "princepadilla@geofarm.com", + "phone": "+1 (843) 428-3506", + "address": "501 Dunham Place, Beaulieu, Utah, 9004", + "registered": "2016-05-01T11:08:29 -02:00", + "latitude": -58.751416, + "longitude": -135.784883, + "tags": [ + "id", + "tempor", + "irure", + "veniam", + "in", + "sit", + "pariatur" + ], + "friends": [ + { + "id": 0, + "name": "Lou Hughes" + }, + { + "id": 1, + "name": "Kayla Chavez" + }, + { + "id": 2, + "name": "Underwood Schmidt" + } + ], + "greeting": "Hello, Prince Padilla! You have 2 unread messages.", + "favoriteFruit": "strawberry" + }, + { + "_id": "58fa646df761ea7972ffdc5e", + "index": 211, + "guid": "3c638abf-110c-450c-ad41-37ebe84ee1c8", + "isActive": true, + "balance": "$3,276.29", + "picture": "http://placehold.it/32x32", + "age": 28, + "eyeColor": "green", + "name": "Hilda Hahn", + "gender": "female", + "company": "HYPLEX", + "email": "hildahahn@hyplex.com", + "phone": "+1 (943) 465-2609", + "address": "348 Crescent Street, Oceola, Georgia, 2167", + "registered": "2015-04-21T11:09:23 -02:00", + "latitude": -64.971497, + "longitude": 98.879269, + "tags": [ + "proident", + "culpa", + "pariatur", + "enim", + "occaecat", + "et", + "ullamco" + ], + "friends": [ + { + "id": 0, + "name": "Merle Berger" + }, + { + "id": 1, + "name": "Alston Higgins" + }, + { + "id": 2, + "name": "Thornton Morton" + } + ], + "greeting": "Hello, Hilda Hahn! You have 2 unread messages.", + "favoriteFruit": "apple" + }, + { + "_id": "58fa646d8ccb02186fdcde3f", + "index": 212, + "guid": "f16e1f47-8f60-4922-9c40-67dd8f8509b1", + "isActive": true, + "balance": "$1,334.71", + "picture": "http://placehold.it/32x32", + "age": 40, + "eyeColor": "brown", + "name": "Dillard Lee", + "gender": "male", + "company": "NAMEBOX", + "email": "dillardlee@namebox.com", + "phone": "+1 (945) 590-2954", + "address": "353 Taaffe Place, Maury, Pennsylvania, 601", + "registered": "2017-04-07T04:56:45 -02:00", + "latitude": 58.646932, + "longitude": -101.897824, + "tags": [ + "aute", + "ex", + "laborum", + "veniam", + "labore", + "nisi", + "dolor" + ], + "friends": [ + { + "id": 0, + "name": "Maricela Rice" + }, + { + "id": 1, + "name": "Workman Cole" + }, + { + "id": 2, + "name": "Debra Robinson" + } + ], + "greeting": "Hello, Dillard Lee! You have 2 unread messages.", + "favoriteFruit": "strawberry" + }, + { + "_id": "58fa646d57b8c10eba7ffa73", + "index": 213, + "guid": "1b3324d4-92d4-4f56-bb46-337ea4968ac4", + "isActive": false, + "balance": "$1,777.43", + "picture": "http://placehold.it/32x32", + "age": 30, + "eyeColor": "green", + "name": "Atkinson Barlow", + "gender": "male", + "company": "FORTEAN", + "email": "atkinsonbarlow@fortean.com", + "phone": "+1 (980) 512-2958", + "address": "773 Girard Street, Hatteras, Indiana, 7062", + "registered": "2017-02-07T12:32:14 -01:00", + "latitude": 46.351996, + "longitude": 73.524112, + "tags": [ + "nisi", + "in", + "do", + "pariatur", + "in", + "magna", + "do" + ], + "friends": [ + { + "id": 0, + "name": "Spencer Palmer" + }, + { + "id": 1, + "name": "Bridgett Aguilar" + }, + { + "id": 2, + "name": "Petty David" + } + ], + "greeting": "Hello, Atkinson Barlow! You have 1 unread messages.", + "favoriteFruit": "strawberry" + }, + { + "_id": "58fa646d93e0d980d0c53fd9", + "index": 214, + "guid": "3ee06883-3f20-49be-ae5d-ee8837523323", + "isActive": true, + "balance": "$3,289.46", + "picture": "http://placehold.it/32x32", + "age": 38, + "eyeColor": "green", + "name": "Nguyen Barron", + "gender": "male", + "company": "VORTEXACO", + "email": "nguyenbarron@vortexaco.com", + "phone": "+1 (862) 426-3930", + "address": "834 Veranda Place, Edgewater, District Of Columbia, 1407", + "registered": "2014-03-26T05:45:51 -01:00", + "latitude": -21.256991, + "longitude": 123.984499, + "tags": [ + "nostrud", + "amet", + "dolore", + "adipisicing", + "anim", + "culpa", + "sit" + ], + "friends": [ + { + "id": 0, + "name": "Deann Alvarado" + }, + { + "id": 1, + "name": "Lott Graham" + }, + { + "id": 2, + "name": "Deirdre Kent" + } + ], + "greeting": "Hello, Nguyen Barron! You have 9 unread messages.", + "favoriteFruit": "banana" + }, + { + "_id": "58fa646d38181a919498a8cb", + "index": 215, + "guid": "f1d57206-abac-48bf-8e2f-e497790e7184", + "isActive": false, + "balance": "$3,035.52", + "picture": "http://placehold.it/32x32", + "age": 20, + "eyeColor": "blue", + "name": "Monique Whitney", + "gender": "female", + "company": "SKYPLEX", + "email": "moniquewhitney@skyplex.com", + "phone": "+1 (841) 524-2218", + "address": "556 Stratford Road, Chicopee, South Dakota, 9508", + "registered": "2016-08-26T11:06:24 -02:00", + "latitude": 43.673768, + "longitude": 66.988309, + "tags": [ + "irure", + "aliqua", + "irure", + "aute", + "consequat", + "esse", + "officia" + ], + "friends": [ + { + "id": 0, + "name": "Battle Roth" + }, + { + "id": 1, + "name": "Cecile Sutton" + }, + { + "id": 2, + "name": "Katy Marquez" + } + ], + "greeting": "Hello, Monique Whitney! You have 9 unread messages.", + "favoriteFruit": "banana" + }, + { + "_id": "58fa646dcf1f8d747bfc1cc5", + "index": 216, + "guid": "4cc2dab6-ab92-469a-8ee8-bb276d6988b5", + "isActive": false, + "balance": "$3,661.14", + "picture": "http://placehold.it/32x32", + "age": 36, + "eyeColor": "brown", + "name": "Preston Bright", + "gender": "male", + "company": "QIMONK", + "email": "prestonbright@qimonk.com", + "phone": "+1 (894) 590-3856", + "address": "205 Highlawn Avenue, Edinburg, Ohio, 4065", + "registered": "2014-05-15T12:31:24 -02:00", + "latitude": -28.515367, + "longitude": -107.312396, + "tags": [ + "dolor", + "veniam", + "non", + "voluptate", + "culpa", + "in", + "do" + ], + "friends": [ + { + "id": 0, + "name": "Pratt Wheeler" + }, + { + "id": 1, + "name": "Alisha Delgado" + }, + { + "id": 2, + "name": "Stevens Shaffer" + } + ], + "greeting": "Hello, Preston Bright! You have 7 unread messages.", + "favoriteFruit": "banana" + }, + { + "_id": "58fa646dc0a993f54f1593f7", + "index": 217, + "guid": "5ed396bf-2287-4026-91d7-538288e13269", + "isActive": true, + "balance": "$1,790.48", + "picture": "http://placehold.it/32x32", + "age": 30, + "eyeColor": "green", + "name": "Mullen Mayo", + "gender": "male", + "company": "SLAMBDA", + "email": "mullenmayo@slambda.com", + "phone": "+1 (879) 459-2741", + "address": "698 Louisa Street, Topaz, Palau, 2158", + "registered": "2016-06-20T11:58:44 -02:00", + "latitude": 79.659776, + "longitude": 17.069697, + "tags": [ + "ad", + "nisi", + "incididunt", + "sint", + "exercitation", + "eiusmod", + "duis" + ], + "friends": [ + { + "id": 0, + "name": "Sasha Pittman" + }, + { + "id": 1, + "name": "Townsend Wiggins" + }, + { + "id": 2, + "name": "Abigail Blankenship" + } + ], + "greeting": "Hello, Mullen Mayo! You have 8 unread messages.", + "favoriteFruit": "strawberry" + }, + { + "_id": "58fa646d197e84ac566aecce", + "index": 218, + "guid": "128f74e2-4570-443c-a7dd-b3ee51c1f8b8", + "isActive": true, + "balance": "$1,030.61", + "picture": "http://placehold.it/32x32", + "age": 39, + "eyeColor": "green", + "name": "Blair Holman", + "gender": "male", + "company": "ZYTREK", + "email": "blairholman@zytrek.com", + "phone": "+1 (801) 505-2127", + "address": "561 Pierrepont Place, Noxen, Louisiana, 2976", + "registered": "2017-02-17T08:48:34 -01:00", + "latitude": 5.585257, + "longitude": -95.035584, + "tags": [ + "adipisicing", + "labore", + "consectetur", + "reprehenderit", + "nulla", + "minim", + "nisi" + ], + "friends": [ + { + "id": 0, + "name": "Frederick Rosario" + }, + { + "id": 1, + "name": "Kane Gutierrez" + }, + { + "id": 2, + "name": "Maynard Duffy" + } + ], + "greeting": "Hello, Blair Holman! You have 2 unread messages.", + "favoriteFruit": "strawberry" + }, + { + "_id": "58fa646d1f728a68d2f3417f", + "index": 219, + "guid": "5297f259-2165-4edd-bfa9-2660c96a4356", + "isActive": true, + "balance": "$1,892.17", + "picture": "http://placehold.it/32x32", + "age": 20, + "eyeColor": "brown", + "name": "Lakeisha Good", + "gender": "female", + "company": "ISOLOGICA", + "email": "lakeishagood@isologica.com", + "phone": "+1 (931) 456-3941", + "address": "615 Putnam Avenue, Longoria, Montana, 5039", + "registered": "2015-07-31T05:37:08 -02:00", + "latitude": 69.711024, + "longitude": -124.480195, + "tags": [ + "amet", + "exercitation", + "ad", + "officia", + "ex", + "minim", + "elit" + ], + "friends": [ + { + "id": 0, + "name": "Veronica Fuller" + }, + { + "id": 1, + "name": "Latisha Meyers" + }, + { + "id": 2, + "name": "Eliza Mcclure" + } + ], + "greeting": "Hello, Lakeisha Good! You have 10 unread messages.", + "favoriteFruit": "banana" + }, + { + "_id": "58fa646db0dca5ff442dd9d4", + "index": 220, + "guid": "062bc6e7-9198-49cb-baeb-ab4887b4fa21", + "isActive": true, + "balance": "$3,074.81", + "picture": "http://placehold.it/32x32", + "age": 22, + "eyeColor": "brown", + "name": "Morrison Norton", + "gender": "male", + "company": "TELLIFLY", + "email": "morrisonnorton@tellifly.com", + "phone": "+1 (834) 457-3879", + "address": "978 Harwood Place, Roosevelt, Illinois, 3081", + "registered": "2014-08-14T10:06:42 -02:00", + "latitude": 27.668418, + "longitude": -79.16662, + "tags": [ + "sunt", + "velit", + "veniam", + "ea", + "tempor", + "ea", + "tempor" + ], + "friends": [ + { + "id": 0, + "name": "Leah Kline" + }, + { + "id": 1, + "name": "Odom Herman" + }, + { + "id": 2, + "name": "Emma Gentry" + } + ], + "greeting": "Hello, Morrison Norton! You have 9 unread messages.", + "favoriteFruit": "banana" + }, + { + "_id": "58fa646dd968efe4033cc90c", + "index": 221, + "guid": "2f1f6c43-9fa5-43dc-9af0-5a692d03f049", + "isActive": true, + "balance": "$3,460.66", + "picture": "http://placehold.it/32x32", + "age": 38, + "eyeColor": "blue", + "name": "Horton Morales", + "gender": "male", + "company": "ZANITY", + "email": "hortonmorales@zanity.com", + "phone": "+1 (990) 471-2505", + "address": "939 Knickerbocker Avenue, Rodanthe, Minnesota, 1001", + "registered": "2016-02-22T12:40:54 -01:00", + "latitude": 58.694956, + "longitude": 116.664682, + "tags": [ + "officia", + "irure", + "sint", + "reprehenderit", + "incididunt", + "aliqua", + "adipisicing" + ], + "friends": [ + { + "id": 0, + "name": "Virgie Short" + }, + { + "id": 1, + "name": "Ladonna Mueller" + }, + { + "id": 2, + "name": "Emerson Lynn" + } + ], + "greeting": "Hello, Horton Morales! You have 5 unread messages.", + "favoriteFruit": "banana" + }, + { + "_id": "58fa646db64960ab412393ac", + "index": 222, + "guid": "8635596a-013a-432b-987d-d3473d9bc320", + "isActive": false, + "balance": "$2,600.62", + "picture": "http://placehold.it/32x32", + "age": 37, + "eyeColor": "blue", + "name": "Margie Hartman", + "gender": "female", + "company": "BRISTO", + "email": "margiehartman@bristo.com", + "phone": "+1 (973) 517-2539", + "address": "736 Cropsey Avenue, Tedrow, South Carolina, 730", + "registered": "2016-08-05T09:03:26 -02:00", + "latitude": -89.862078, + "longitude": 60.151029, + "tags": [ + "occaecat", + "elit", + "fugiat", + "pariatur", + "esse", + "velit", + "anim" + ], + "friends": [ + { + "id": 0, + "name": "Marshall Cook" + }, + { + "id": 1, + "name": "Gilmore Velez" + }, + { + "id": 2, + "name": "Hollie Garner" + } + ], + "greeting": "Hello, Margie Hartman! You have 9 unread messages.", + "favoriteFruit": "banana" + }, + { + "_id": "58fa646d40b20161724b84dd", + "index": 223, + "guid": "1741c2fa-9376-4181-b3d0-d1c6f8289272", + "isActive": false, + "balance": "$1,194.01", + "picture": "http://placehold.it/32x32", + "age": 33, + "eyeColor": "blue", + "name": "Nadine Ochoa", + "gender": "female", + "company": "HOTCAKES", + "email": "nadineochoa@hotcakes.com", + "phone": "+1 (856) 479-3353", + "address": "758 Dunne Place, Breinigsville, New Mexico, 8595", + "registered": "2017-04-08T02:57:06 -02:00", + "latitude": -23.177532, + "longitude": -128.194926, + "tags": [ + "non", + "commodo", + "proident", + "voluptate", + "duis", + "officia", + "laborum" + ], + "friends": [ + { + "id": 0, + "name": "Leticia Macias" + }, + { + "id": 1, + "name": "Robyn Hampton" + }, + { + "id": 2, + "name": "Barker Watts" + } + ], + "greeting": "Hello, Nadine Ochoa! You have 10 unread messages.", + "favoriteFruit": "apple" + }, + { + "_id": "58fa646d468e1d0e28a9a98c", + "index": 224, + "guid": "f1be88ee-67a4-49fe-95a7-d744d348b893", + "isActive": false, + "balance": "$1,808.05", + "picture": "http://placehold.it/32x32", + "age": 35, + "eyeColor": "brown", + "name": "Lauren Fulton", + "gender": "female", + "company": "SKINSERVE", + "email": "laurenfulton@skinserve.com", + "phone": "+1 (886) 599-3798", + "address": "333 Cass Place, Morgandale, North Carolina, 9217", + "registered": "2017-03-25T06:42:38 -01:00", + "latitude": -88.867505, + "longitude": -156.964591, + "tags": [ + "voluptate", + "sit", + "sit", + "amet", + "nisi", + "deserunt", + "Lorem" + ], + "friends": [ + { + "id": 0, + "name": "Hooper Mckay" + }, + { + "id": 1, + "name": "Rose Peck" + }, + { + "id": 2, + "name": "Heather Osborn" + } + ], + "greeting": "Hello, Lauren Fulton! You have 8 unread messages.", + "favoriteFruit": "strawberry" + }, + { + "_id": "58fa646d34d32d8c9d2b6ac0", + "index": 225, + "guid": "b7592a1d-cb81-4cc8-ac6b-aa430880728a", + "isActive": false, + "balance": "$1,483.97", + "picture": "http://placehold.it/32x32", + "age": 26, + "eyeColor": "brown", + "name": "Josefa Hines", + "gender": "female", + "company": "OTHERSIDE", + "email": "josefahines@otherside.com", + "phone": "+1 (878) 537-3702", + "address": "600 Harbor Lane, Chesapeake, Massachusetts, 6097", + "registered": "2016-01-10T01:21:51 -01:00", + "latitude": 9.11098, + "longitude": -17.908215, + "tags": [ + "est", + "laborum", + "amet", + "aute", + "sunt", + "exercitation", + "ea" + ], + "friends": [ + { + "id": 0, + "name": "Ophelia Sandoval" + }, + { + "id": 1, + "name": "Bridget Shannon" + }, + { + "id": 2, + "name": "Estella Townsend" + } + ], + "greeting": "Hello, Josefa Hines! You have 8 unread messages.", + "favoriteFruit": "strawberry" + }, + { + "_id": "58fa646d74716e57363fd23f", + "index": 226, + "guid": "4b2ebddb-e23f-4c31-8dd5-6686bbdd28e8", + "isActive": true, + "balance": "$1,137.98", + "picture": "http://placehold.it/32x32", + "age": 29, + "eyeColor": "green", + "name": "Reva Galloway", + "gender": "female", + "company": "XSPORTS", + "email": "revagalloway@xsports.com", + "phone": "+1 (884) 463-2654", + "address": "970 Elton Street, Winesburg, Kansas, 8469", + "registered": "2014-04-12T03:14:08 -02:00", + "latitude": -61.152244, + "longitude": -134.701953, + "tags": [ + "qui", + "qui", + "commodo", + "fugiat", + "sit", + "aliqua", + "qui" + ], + "friends": [ + { + "id": 0, + "name": "Sonia Delacruz" + }, + { + "id": 1, + "name": "Tasha Matthews" + }, + { + "id": 2, + "name": "Andrews Green" + } + ], + "greeting": "Hello, Reva Galloway! You have 4 unread messages.", + "favoriteFruit": "strawberry" + }, + { + "_id": "58fa646de8b8f2fc424d77e6", + "index": 227, + "guid": "42e7f9f2-16a9-4a72-995f-135a9cc5d46f", + "isActive": true, + "balance": "$3,677.44", + "picture": "http://placehold.it/32x32", + "age": 30, + "eyeColor": "green", + "name": "Brooke Lopez", + "gender": "female", + "company": "ARCHITAX", + "email": "brookelopez@architax.com", + "phone": "+1 (964) 562-3013", + "address": "938 Strong Place, Dante, American Samoa, 1982", + "registered": "2016-10-01T08:24:31 -02:00", + "latitude": 84.487681, + "longitude": 86.796642, + "tags": [ + "et", + "irure", + "anim", + "deserunt", + "elit", + "velit", + "commodo" + ], + "friends": [ + { + "id": 0, + "name": "Pena Lambert" + }, + { + "id": 1, + "name": "Patrice Guthrie" + }, + { + "id": 2, + "name": "Miranda Ross" + } + ], + "greeting": "Hello, Brooke Lopez! You have 6 unread messages.", + "favoriteFruit": "strawberry" + }, + { + "_id": "58fa646d2f9a3877a506b0e1", + "index": 228, + "guid": "99dd12d8-576f-4f85-a12a-1ed3e93d598d", + "isActive": false, + "balance": "$3,642.51", + "picture": "http://placehold.it/32x32", + "age": 33, + "eyeColor": "blue", + "name": "Belinda Gamble", + "gender": "female", + "company": "COMBOGENE", + "email": "belindagamble@combogene.com", + "phone": "+1 (906) 416-3652", + "address": "568 Holt Court, Blue, Tennessee, 3439", + "registered": "2016-12-02T04:22:40 -01:00", + "latitude": -88.861634, + "longitude": 100.548692, + "tags": [ + "velit", + "consequat", + "dolore", + "aliqua", + "qui", + "enim", + "qui" + ], + "friends": [ + { + "id": 0, + "name": "Price Winters" + }, + { + "id": 1, + "name": "Earline Hancock" + }, + { + "id": 2, + "name": "Vincent Brennan" + } + ], + "greeting": "Hello, Belinda Gamble! You have 8 unread messages.", + "favoriteFruit": "strawberry" + }, + { + "_id": "58fa646d202ac7bb4df7bb98", + "index": 229, + "guid": "93465bba-20ba-4019-bfb4-e8500d4fb217", + "isActive": true, + "balance": "$1,083.71", + "picture": "http://placehold.it/32x32", + "age": 28, + "eyeColor": "green", + "name": "Ernestine Hayes", + "gender": "female", + "company": "UTARA", + "email": "ernestinehayes@utara.com", + "phone": "+1 (884) 522-3278", + "address": "694 Alton Place, Matheny, Florida, 6179", + "registered": "2015-04-12T09:44:31 -02:00", + "latitude": -0.87388, + "longitude": -76.646898, + "tags": [ + "tempor", + "irure", + "pariatur", + "consequat", + "sint", + "reprehenderit", + "aliqua" + ], + "friends": [ + { + "id": 0, + "name": "Small Deleon" + }, + { + "id": 1, + "name": "Gayle Moss" + }, + { + "id": 2, + "name": "Mack Ortega" + } + ], + "greeting": "Hello, Ernestine Hayes! You have 8 unread messages.", + "favoriteFruit": "banana" + }, + { + "_id": "58fa646d70a138930a40e586", + "index": 230, + "guid": "775309a2-30e1-465c-bcb1-c04a1f7de906", + "isActive": true, + "balance": "$3,665.22", + "picture": "http://placehold.it/32x32", + "age": 26, + "eyeColor": "brown", + "name": "Bertha Berg", + "gender": "female", + "company": "ACCEL", + "email": "berthaberg@accel.com", + "phone": "+1 (952) 594-2803", + "address": "175 Hoyt Street, Neibert, Missouri, 1361", + "registered": "2014-11-27T04:31:37 -01:00", + "latitude": 86.737132, + "longitude": 43.892186, + "tags": [ + "et", + "laboris", + "irure", + "duis", + "anim", + "aliquip", + "dolore" + ], + "friends": [ + { + "id": 0, + "name": "Nichols House" + }, + { + "id": 1, + "name": "Jerry Bradley" + }, + { + "id": 2, + "name": "Dale Collins" + } + ], + "greeting": "Hello, Bertha Berg! You have 3 unread messages.", + "favoriteFruit": "apple" + }, + { + "_id": "58fa646d2f6f7ce0156d4704", + "index": 231, + "guid": "a29b3093-5430-4ceb-933e-0959960260d6", + "isActive": true, + "balance": "$1,605.78", + "picture": "http://placehold.it/32x32", + "age": 28, + "eyeColor": "brown", + "name": "Carla Caldwell", + "gender": "female", + "company": "BUZZMAKER", + "email": "carlacaldwell@buzzmaker.com", + "phone": "+1 (863) 465-3689", + "address": "974 Lloyd Street, Fedora, Marshall Islands, 6289", + "registered": "2016-08-22T02:50:55 -02:00", + "latitude": 32.202462, + "longitude": 176.179413, + "tags": [ + "laborum", + "elit", + "occaecat", + "Lorem", + "culpa", + "laboris", + "aliquip" + ], + "friends": [ + { + "id": 0, + "name": "Rosanna Robbins" + }, + { + "id": 1, + "name": "Marci Hutchinson" + }, + { + "id": 2, + "name": "Brown Mclaughlin" + } + ], + "greeting": "Hello, Carla Caldwell! You have 1 unread messages.", + "favoriteFruit": "banana" + }, + { + "_id": "58fa646d5218753d8c52a7b5", + "index": 232, + "guid": "f35586ff-f165-43c2-8465-da96448e40a0", + "isActive": true, + "balance": "$2,764.50", + "picture": "http://placehold.it/32x32", + "age": 29, + "eyeColor": "green", + "name": "Justice Burris", + "gender": "male", + "company": "FLUMBO", + "email": "justiceburris@flumbo.com", + "phone": "+1 (876) 579-3387", + "address": "206 Love Lane, Whitewater, Washington, 3072", + "registered": "2016-09-01T08:05:38 -02:00", + "latitude": 71.16566, + "longitude": 69.95058, + "tags": [ + "voluptate", + "ipsum", + "consequat", + "aliquip", + "dolore", + "aute", + "quis" + ], + "friends": [ + { + "id": 0, + "name": "Isabelle Shelton" + }, + { + "id": 1, + "name": "Tiffany Farrell" + }, + { + "id": 2, + "name": "Banks Moon" + } + ], + "greeting": "Hello, Justice Burris! You have 10 unread messages.", + "favoriteFruit": "banana" + }, + { + "_id": "58fa646dad8a49ace74d2649", + "index": 233, + "guid": "e5c46f5d-927e-4bbe-90c3-cf8bbe26a68f", + "isActive": false, + "balance": "$3,862.66", + "picture": "http://placehold.it/32x32", + "age": 25, + "eyeColor": "brown", + "name": "Yvette Petty", + "gender": "female", + "company": "LEXICONDO", + "email": "yvettepetty@lexicondo.com", + "phone": "+1 (939) 479-3944", + "address": "464 Lois Avenue, Ribera, Colorado, 8527", + "registered": "2016-05-02T01:08:18 -02:00", + "latitude": 17.151992, + "longitude": 10.854582, + "tags": [ + "do", + "enim", + "cillum", + "excepteur", + "Lorem", + "minim", + "irure" + ], + "friends": [ + { + "id": 0, + "name": "Mable Olson" + }, + { + "id": 1, + "name": "Love Hamilton" + }, + { + "id": 2, + "name": "Neal Sears" + } + ], + "greeting": "Hello, Yvette Petty! You have 4 unread messages.", + "favoriteFruit": "banana" + }, + { + "_id": "58fa646d23c48f21434c1d71", + "index": 234, + "guid": "8596835f-7b66-4948-a5b7-ed98653ba885", + "isActive": true, + "balance": "$1,455.28", + "picture": "http://placehold.it/32x32", + "age": 37, + "eyeColor": "brown", + "name": "Dudley Kinney", + "gender": "male", + "company": "XOGGLE", + "email": "dudleykinney@xoggle.com", + "phone": "+1 (976) 573-3513", + "address": "245 Nelson Street, Homestead, Puerto Rico, 6198", + "registered": "2016-12-29T01:22:08 -01:00", + "latitude": -85.068498, + "longitude": -80.978942, + "tags": [ + "incididunt", + "enim", + "veniam", + "qui", + "culpa", + "dolore", + "ipsum" + ], + "friends": [ + { + "id": 0, + "name": "Erin Carr" + }, + { + "id": 1, + "name": "Caldwell Miller" + }, + { + "id": 2, + "name": "Lorene Hicks" + } + ], + "greeting": "Hello, Dudley Kinney! You have 1 unread messages.", + "favoriteFruit": "banana" + }, + { + "_id": "58fa646d4cae4bf832b10a78", + "index": 235, + "guid": "e31f6d5e-792c-4663-8a03-d0bdd6b27174", + "isActive": false, + "balance": "$1,902.69", + "picture": "http://placehold.it/32x32", + "age": 31, + "eyeColor": "blue", + "name": "Cara Sims", + "gender": "female", + "company": "SEQUITUR", + "email": "carasims@sequitur.com", + "phone": "+1 (922) 413-2438", + "address": "301 Arlington Place, Machias, Rhode Island, 2796", + "registered": "2016-11-25T11:52:04 -01:00", + "latitude": 81.764319, + "longitude": 24.388516, + "tags": [ + "sint", + "tempor", + "do", + "dolore", + "ea", + "excepteur", + "elit" + ], + "friends": [ + { + "id": 0, + "name": "Castro Roman" + }, + { + "id": 1, + "name": "Celia Hensley" + }, + { + "id": 2, + "name": "Jana Conley" + } + ], + "greeting": "Hello, Cara Sims! You have 2 unread messages.", + "favoriteFruit": "banana" + }, + { + "_id": "58fa646d02ad270f8e98f9df", + "index": 236, + "guid": "c94d4b73-bd68-43e7-a2b5-87f3fa14e0bb", + "isActive": false, + "balance": "$3,825.22", + "picture": "http://placehold.it/32x32", + "age": 40, + "eyeColor": "blue", + "name": "Levine Schneider", + "gender": "male", + "company": "KEGULAR", + "email": "levineschneider@kegular.com", + "phone": "+1 (976) 481-3448", + "address": "807 Meeker Avenue, Rutherford, Alabama, 8525", + "registered": "2015-07-02T10:10:12 -02:00", + "latitude": 64.409918, + "longitude": 112.027182, + "tags": [ + "aliqua", + "nostrud", + "proident", + "cillum", + "magna", + "non", + "duis" + ], + "friends": [ + { + "id": 0, + "name": "Blanche Hopper" + }, + { + "id": 1, + "name": "Jacqueline Camacho" + }, + { + "id": 2, + "name": "Sue Huffman" + } + ], + "greeting": "Hello, Levine Schneider! You have 6 unread messages.", + "favoriteFruit": "apple" + }, + { + "_id": "58fa646d56246e5eff4b4373", + "index": 237, + "guid": "614791e1-b12d-4ac7-9a82-e2f44df3d60c", + "isActive": true, + "balance": "$2,564.32", + "picture": "http://placehold.it/32x32", + "age": 28, + "eyeColor": "brown", + "name": "Eddie Dillard", + "gender": "female", + "company": "LETPRO", + "email": "eddiedillard@letpro.com", + "phone": "+1 (997) 438-2314", + "address": "514 Suydam Place, Enetai, Wisconsin, 8450", + "registered": "2017-04-14T12:09:22 -02:00", + "latitude": 46.420077, + "longitude": -157.287046, + "tags": [ + "voluptate", + "sunt", + "veniam", + "nulla", + "eu", + "cillum", + "sint" + ], + "friends": [ + { + "id": 0, + "name": "Norris Sosa" + }, + { + "id": 1, + "name": "Darla Whitley" + }, + { + "id": 2, + "name": "Marlene Bolton" + } + ], + "greeting": "Hello, Eddie Dillard! You have 9 unread messages.", + "favoriteFruit": "strawberry" + }, + { + "_id": "58fa646d60f0a469a2d17bac", + "index": 238, + "guid": "f90329fd-4346-4ff1-a48a-b72b1f12d674", + "isActive": true, + "balance": "$1,182.07", + "picture": "http://placehold.it/32x32", + "age": 28, + "eyeColor": "green", + "name": "Brandi Cannon", + "gender": "female", + "company": "GEEKNET", + "email": "brandicannon@geeknet.com", + "phone": "+1 (888) 580-2037", + "address": "531 Dodworth Street, Layhill, California, 8537", + "registered": "2014-11-02T12:50:59 -01:00", + "latitude": 63.581991, + "longitude": -47.852829, + "tags": [ + "est", + "elit", + "ex", + "aliqua", + "ea", + "in", + "tempor" + ], + "friends": [ + { + "id": 0, + "name": "Cornelia Wilcox" + }, + { + "id": 1, + "name": "Saundra Gomez" + }, + { + "id": 2, + "name": "Terri Fischer" + } + ], + "greeting": "Hello, Brandi Cannon! You have 5 unread messages.", + "favoriteFruit": "banana" + }, + { + "_id": "58fa646d3af6460339584fe7", + "index": 239, + "guid": "a1557310-233d-41d5-9440-55add549f99f", + "isActive": false, + "balance": "$1,217.25", + "picture": "http://placehold.it/32x32", + "age": 34, + "eyeColor": "brown", + "name": "Hilary Wolf", + "gender": "female", + "company": "CALLFLEX", + "email": "hilarywolf@callflex.com", + "phone": "+1 (879) 442-2911", + "address": "868 Dewitt Avenue, Darrtown, Arkansas, 7415", + "registered": "2016-04-05T08:07:20 -02:00", + "latitude": 72.925737, + "longitude": -122.846033, + "tags": [ + "id", + "dolor", + "ex", + "magna", + "nisi", + "consequat", + "pariatur" + ], + "friends": [ + { + "id": 0, + "name": "Cotton Vazquez" + }, + { + "id": 1, + "name": "Jennie Neal" + }, + { + "id": 2, + "name": "Penelope Wright" + } + ], + "greeting": "Hello, Hilary Wolf! You have 8 unread messages.", + "favoriteFruit": "strawberry" + } +] diff --git a/tests/bench/fmtSExprFull.lean b/tests/bench/fmtSExprFull.lean new file mode 100644 index 000000000000..dacb019fc85d --- /dev/null +++ b/tests/bench/fmtSExprFull.lean @@ -0,0 +1,54 @@ +import Lean.Data.Fmt.Formatter + +open Lean.Fmt + +inductive SExpr where + | leaf (v : String) + | node (cs : List SExpr) + +def asConcat (ds : List Doc) : Doc := + match ds with + | [] => .failure + | [d] => d + | d :: ds => + .joinUsing (.text " ") (#[d] ++ ds.toArray.map Doc.aligned) + +def pp (s : SExpr) : Doc := + match s with + | .leaf v => + .text v + | .node cs => + let cs := cs.map pp + .join #[ + .text "(", + .aligned (.either + (.joinUsing .hardNl cs.toArray) + (asConcat cs)), + .aligned (.text ")") + ] + +def testExpr (n c : Nat) : SExpr × Nat := + if n = 0 then + (.leaf (toString c), c + 1) + else + let (t1, c1) := testExpr (n - 1) c + let (t2, c2) := testExpr (n - 1) c1 + (.node [t1, t2], c2) + +@[noinline] +def doc (n : Nat) : IO Doc := + return pp (testExpr n 0).1 + +@[noinline] +def format (doc : Doc) : IO (Option String) := do + return format? doc 80 100 + +def main (args : List String) : IO Unit := do + let n := (args[0]!).toNat! + let d ← doc n + let startNs ← IO.monoNanosNow + let r? ← format d + let endNs ← IO.monoNanosNow + let benchTime : Float := (endNs - startNs).toFloat / 1_000_000_000.0 + assert! r?.isSome + IO.println s!"format: {benchTime}" diff --git a/tests/bench/run_fmt_bench.sh b/tests/bench/run_fmt_bench.sh new file mode 100755 index 000000000000..ce51e1c19eb4 --- /dev/null +++ b/tests/bench/run_fmt_bench.sh @@ -0,0 +1,35 @@ +#!/usr/bin/env bash + +# Usage: +# ./run_all.sh file1 "args for file1" file2 "args for file2" ... + +# Ensure arguments come in pairs +if (( $# % 2 != 0 )); then + echo "Error: arguments must come in pairs." + exit 1 +fi + +# Loop over pairs +while (( "$#" )); do + file="$1" + args="$2" + shift 2 + + echo "Processing file: $file" + echo "Args: $args" + + for i in {1..3}; do + echo "---- Run $i for $file ----" + + # Compile + ./compile.sh "$file" || { + echo "Compilation failed for $file" + break + } + + # Execute + "./${file}.out" $args + done + + echo "-------------------------" +done diff --git a/tests/bench/speedcenter.exec.velcom.yaml b/tests/bench/speedcenter.exec.velcom.yaml index 904895414121..29f73f33488f 100644 --- a/tests/bench/speedcenter.exec.velcom.yaml +++ b/tests/bench/speedcenter.exec.velcom.yaml @@ -295,6 +295,51 @@ <<: *time cmd: lean -Dlinter.all=false --run identifier_completion_runner.lean parse_output: true +- attributes: + description: format (flatten) + tags: [other] + run_config: + <<: *time + cmd: bash -c "ulimit -s unlimited && ./fmtFlatten.lean.out 16000" + parse_output: true + build_config: + cmd: ./compile.sh fmtFlatten.lean +- attributes: + description: format (concat) + tags: [other] + run_config: + <<: *time + cmd: bash -c "ulimit -s unlimited && ./fmtConcat.lean.out 50000" + parse_output: true + build_config: + cmd: ./compile.sh fmtConcat.lean +- attributes: + description: format (sexpr full) + tags: [other] + run_config: + <<: *time + cmd: bash -c "ulimit -s unlimited && ./fmtSExprFull.lean.out 15" + parse_output: true + build_config: + cmd: ./compile.sh fmtSExprFull.lean +- attributes: + description: format (fill-sep) + tags: [other] + run_config: + <<: *time + cmd: bash -c "ulimit -s unlimited && ./fmtFillSep.lean.out 10000" + parse_output: true + build_config: + cmd: ./compile.sh fmtFillSep.lean +- attributes: + description: format (json) + tags: [other] + run_config: + <<: *time + cmd: bash -c "ulimit -s unlimited && ./fmtJson.lean.out" + parse_output: true + build_config: + cmd: ./compile.sh fmtJson.lean - attributes: description: liasolver tags: [other] diff --git a/tests/lean/run/fmt.lean b/tests/lean/run/fmt.lean new file mode 100644 index 000000000000..914e006ff484 --- /dev/null +++ b/tests/lean/run/fmt.lean @@ -0,0 +1,279 @@ +import Lean.Data.Fmt.Formatter + +open Lean.Fmt + +def traditional : Doc := + .join #[ + .text "function append(first,second,third){", + .indented 4 true ( + let f : Doc := .text "first +" + let s : Doc := .text "second +" + let t : Doc := .text "third" + .join #[ + .nl, + .text "return ", + .maybeFlattened (.indented 4 true (.join #[f, .nl, s, .nl, t]))] + ), + .nl, + .text "}" + ] + +def test (d : Doc) (width : Nat) (pre : String := "") : IO Unit := do + IO.println "" + let r := pre ++ (format? d width (offset := pre.length) |>.getD "") + IO.println r + +/-- +info: +function append(first,second,third){ + return first + + second + + third +} +-/ +#guard_msgs in +#eval test traditional 22 + +/-- +info: +function append(first,second,third){ + return first + second + third +} +-/ +#guard_msgs in +#eval test traditional 36 + +inductive SExpr where + | leaf (v : String) + | node (cs : List SExpr) + +instance : Coe String SExpr where + coe v := .leaf v + +partial def SExpr.pretty (s : SExpr) : Doc := + match s with + | .leaf v => .text v + | .node [] => .text "()" + | .node (f :: args) => + let fp := f.pretty + let argsp := args.toArray.map pretty + .oneOf #[ + .join #[ + .text "(", + .aligned (.joinUsing .hardNl (#[fp] ++ argsp)), + .text ")" + ], + .join #[ + .text "(", + .aligned fp, + .text " ", + .aligned (.joinUsing .hardNl argsp), + .text ")" + ], + .flattened ( + .join #[ + .text "(", + .aligned (.joinUsing (.text " ") (#[fp] ++ argsp)), + .text ")" + ] + ) + ] + +partial def SExpr.pretty' (s : SExpr) : Doc := + match s with + | .leaf v => .text v + | .node [] => .text "()" + | .node (f :: args) => + let fp := f.pretty' + let argsp := args.toArray.map pretty' + .oneOf #[ + .join #[ + .text "(", + .aligned (.joinUsing .hardNl (#[fp] ++ argsp)), + .text ")" + ], + .join #[ + .text "(", + .aligned fp, + .text " ", + .aligned (.joinUsing .hardNl argsp), + .text ")" + ], + .join #[ + .text "(", + .aligned (.joinUsing (.text " ") (#[fp] ++ argsp)), + .text ")" + ] + ] + +partial def SExpr.pretty'' (s : SExpr) : Doc := + match s with + | .leaf v => .text v + | .node [] => .text "()" + | .node (f :: args) => + let fp := f.pretty'' + let argsp := args.toArray.map pretty'' + .oneOf #[ + .join #[ + .text "(", + .nested (.joinUsing .hardNl (#[fp] ++ argsp)), + .text ")" + ], + .flattened ( + .join #[ + .text "(", + .nested (.joinUsing (.text " ") (#[fp] ++ argsp)), + .text ")" + ] + ) + ] + +def testSExpr (e : SExpr) (width : Nat) (pre : String := "") : IO Unit := do + test e.pretty width pre + +def testSExpr' (e : SExpr) (width : Nat) : IO Unit := do + test e.pretty' width + +def testSExpr'' (e : SExpr) (width : Nat) (pre : String := "") : IO Unit := do + test e.pretty'' width pre + +def sExpr1 : SExpr := + .node ["+", .node ["foo", "1", "2"], .node ["bar", "2", "3"], .node ["baz", "3", "4"]] + +/-- +info: +(+ (foo 1 2) + (bar 2 3) + (baz 3 4)) +-/ +#guard_msgs in +#eval testSExpr sExpr1 31 + +/-- +info: +(+ (foo 1 2) (bar 2 3) (baz 3 + 4)) +-/ +#guard_msgs in +#eval testSExpr' sExpr1 31 + +def sExpr2 : SExpr := .node ["+", "123", "456", "789"] + + +/-- +info: +(+ 123 456 789) +-/ +#guard_msgs in +#eval testSExpr sExpr2 15 + +/-- +info: +(+ 123 + 456 + 789) +-/ +#guard_msgs in +#eval testSExpr sExpr2 14 + +/-- +info: +(+ + 123 + 456 + 789) +-/ +#guard_msgs in +#eval testSExpr sExpr2 5 + +/-- +info: +(+ + 123 + 456 + 789) +-/ +#guard_msgs in +#eval testSExpr sExpr2 0 + +def sExpr3 : SExpr := + .node ["a", "b", "c", "d"] + +def sExpr4 : SExpr := + .node [sExpr3, sExpr3, sExpr3, sExpr3] + +def sExpr5 : SExpr := + .node [.node ["abcde", sExpr4], .node ["abcdefgh", sExpr4]] + +def pre1 := "hello: " + +/-- +info: +hello: ((abcde ((a b c d) + (a b c d) + (a b c d) + (a b c d))) + (abcdefgh + ((a b c d) + (a b c d) + (a b c d) + (a b c d)))) +-/ +#guard_msgs in +#eval testSExpr sExpr5 (20 + pre1.length) pre1 + +/-- +info: +hello: ((abcde + ((a b c d) + (a b c d) + (a b c d) + (a b c d))) + (abcdefgh + ((a b c d) + (a b c d) + (a b c d) + (a b c d)))) +-/ +#guard_msgs in +#eval testSExpr'' sExpr5 (20 + pre1.length) pre1 + +/-- +info: +abc +def +-/ +#guard_msgs in +#eval test (.indented 4 true (.unindented (.joinUsing .hardNl #[.text "abc", .text "def"]))) 80 + +/-- +info: +abc + def +-/ +#guard_msgs in +#eval test (.indented 4 true (.joinUsing .hardNl #[.text "abc", .text "def"])) 80 + +/-- +info: +something +-/ +#guard_msgs in +#eval test (.either (.joinUsing .hardNl #[.text "abc", .text "def"]) (.text "something")) 80 + +/-- info: none -/ +#guard_msgs in +#eval format? .failure 80 + +def concated (n : Nat) : Doc := + if n = 0 then + .text "" + else + .append (concated (n - 1)) (.text "line") + +/-- +info: +linelinelinelinelinelinelinelinelinelinelinelinelinelinelinelinelinelinelinelinelinelinelinelinelinelinelinelinelinelinelinelinelinelinelinelinelinelinelinelinelinelinelinelinelinelinelinelinelinelinelinelinelinelinelinelinelinelinelinelinelinelinelinelinelinelinelinelinelinelinelinelinelinelinelinelinelinelinelinelinelinelinelinelinelinelinelinelinelinelinelinelinelinelinelinelinelinelinelineline +-/ +#guard_msgs in +#eval test (concated 100) 100